src/HOL/Library/Parity.thy
author urbanc
Fri, 04 Jan 2008 16:35:22 +0100
changeset 25832 41a014cc44c0
parent 25691 8f8d83af100a
child 25875 536dfdc25e0a
permissions -rw-r--r--
partially adapted to new inversion rules
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
21263
wenzelm
parents: 21256
diff changeset
     1
(*  Title:      HOL/Library/Parity.thy
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
     2
    ID:         $Id$
25600
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
     3
    Author:     Jeremy Avigad, 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
25691
8f8d83af100a switched from PreList to ATP_Linkup
haftmann
parents: 25600
diff changeset
     9
imports ATP_Linkup
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
22473
753123c89d72 explizit "type" superclass
haftmann
parents: 22390
diff changeset
    12
class even_odd = type + 
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
25571
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25502
diff changeset
    19
instantiation int and nat :: even_odd
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
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    32
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    33
subsection {* Even and odd are mutually exclusive *}
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    34
21263
wenzelm
parents: 21256
diff changeset
    35
lemma int_pos_lt_two_imp_zero_or_one:
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    36
    "0 <= x ==> (x::int) < 2 ==> x = 0 | x = 1"
23522
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
    37
  by presburger
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    38
23522
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
    39
lemma neq_one_mod_two [simp, presburger]: 
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
    40
  "((x::int) mod 2 ~= 0) = (x mod 2 = 1)" by presburger
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    41
25600
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
    42
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    43
subsection {* Behavior under integer arithmetic operations *}
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    44
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    45
lemma even_times_anything: "even (x::int) ==> even (x * y)"
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    46
  by (simp add: even_def zmod_zmult1_eq')
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    47
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    48
lemma anything_times_even: "even (y::int) ==> even (x * y)"
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    49
  by (simp add: even_def zmod_zmult1_eq)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    50
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    51
lemma odd_times_odd: "odd (x::int) ==> odd y ==> odd (x * y)"
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    52
  by (simp add: even_def zmod_zmult1_eq)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    53
23522
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
    54
lemma even_product[presburger]: "even((x::int) * y) = (even x | even y)"
21263
wenzelm
parents: 21256
diff changeset
    55
  apply (auto simp add: even_times_anything anything_times_even)
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    56
  apply (rule ccontr)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    57
  apply (auto simp add: odd_times_odd)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    58
  done
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    59
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    60
lemma even_plus_even: "even (x::int) ==> even y ==> even (x + y)"
23522
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
    61
  by presburger
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    62
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    63
lemma even_plus_odd: "even (x::int) ==> odd y ==> odd (x + y)"
23522
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
    64
  by presburger
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    65
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    66
lemma odd_plus_even: "odd (x::int) ==> even y ==> odd (x + y)"
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 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
    70
23522
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
    71
lemma even_sum[presburger]: "even ((x::int) + y) = ((even x & even y) | (odd x & odd y))"
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
    72
  by presburger
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    73
23522
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
    74
lemma even_neg[presburger]: "even (-(x::int)) = even x" by presburger
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    75
21263
wenzelm
parents: 21256
diff changeset
    76
lemma even_difference:
23522
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
    77
    "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
    78
21263
wenzelm
parents: 21256
diff changeset
    79
lemma even_pow_gt_zero:
wenzelm
parents: 21256
diff changeset
    80
    "even (x::int) ==> 0 < n ==> even (x^n)"
wenzelm
parents: 21256
diff changeset
    81
  by (induct n) (auto simp add: even_product)
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    82
23522
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
    83
lemma odd_pow_iff[presburger]: "odd ((x::int) ^ n) \<longleftrightarrow> (n = 0 \<or> odd x)"
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
    84
  apply (induct n, simp_all)
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
    85
  apply presburger
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
    86
  apply (case_tac n, auto)
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
    87
  apply (simp_all add: even_product)
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    88
  done
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    89
23522
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
    90
lemma odd_pow: "odd x ==> odd((x::int)^n)" by (simp add: odd_pow_iff)
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
    91
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
    92
lemma even_power[presburger]: "even ((x::int)^n) = (even x & 0 < n)"
21263
wenzelm
parents: 21256
diff changeset
    93
  apply (auto simp add: even_pow_gt_zero)
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    94
  apply (erule contrapos_pp, erule odd_pow)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    95
  apply (erule contrapos_pp, simp add: even_def)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    96
  done
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    97
23522
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
    98
lemma even_zero[presburger]: "even (0::int)" by presburger
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    99
23522
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
   100
lemma odd_one[presburger]: "odd (1::int)" by presburger
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   101
21263
wenzelm
parents: 21256
diff changeset
   102
lemmas even_odd_simps [simp] = even_def[of "number_of v",standard] even_zero
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   103
  odd_one even_product even_sum even_neg even_difference even_power
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   104
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   105
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   106
subsection {* Equivalent definitions *}
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   107
23522
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
   108
lemma two_times_even_div_two: "even (x::int) ==> 2 * (x div 2) = x" 
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
   109
  by presburger
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   110
21263
wenzelm
parents: 21256
diff changeset
   111
lemma two_times_odd_div_two_plus_one: "odd (x::int) ==>
23522
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
   112
    2 * (x div 2) + 1 = x" by presburger
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   113
23522
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
   114
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
   115
23522
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
   116
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
   117
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   118
subsection {* even and odd for nats *}
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
lemma pos_int_even_equiv_nat_even: "0 \<le> x ==> even x = even (nat x)"
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   121
  by (simp add: even_nat_def)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   122
23522
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
   123
lemma even_nat_product[presburger]: "even((x::nat) * y) = (even x | even y)"
23431
25ca91279a9b change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
huffman
parents: 23309
diff changeset
   124
  by (simp add: even_nat_def int_mult)
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   125
23522
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
   126
lemma even_nat_sum[presburger]: "even ((x::nat) + y) =
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
   127
    ((even x & even y) | (odd x & odd y))" by presburger
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   128
23522
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
   129
lemma even_nat_difference[presburger]:
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   130
    "even ((x::nat) - y) = (x < y | (even x & even y) | (odd x & odd y))"
23522
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
   131
by presburger
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   132
23522
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
   133
lemma even_nat_Suc[presburger]: "even (Suc x) = odd x" by presburger
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   134
23522
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
   135
lemma even_nat_power[presburger]: "even ((x::nat)^y) = (even x & 0 < y)"
23431
25ca91279a9b change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
huffman
parents: 23309
diff changeset
   136
  by (simp add: even_nat_def int_power)
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   137
23522
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
   138
lemma even_nat_zero[presburger]: "even (0::nat)" by presburger
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   139
21263
wenzelm
parents: 21256
diff changeset
   140
lemmas even_odd_nat_simps [simp] = even_nat_def[of "number_of v",standard]
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   141
  even_nat_zero even_nat_Suc even_nat_product even_nat_sum even_nat_power
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   142
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   143
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   144
subsection {* Equivalent definitions *}
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   145
21263
wenzelm
parents: 21256
diff changeset
   146
lemma nat_lt_two_imp_zero_or_one: "(x::nat) < Suc (Suc 0) ==>
23522
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
   147
    x = 0 | x = Suc 0" by presburger
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   148
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   149
lemma even_nat_mod_two_eq_zero: "even (x::nat) ==> x mod (Suc (Suc 0)) = 0"
23522
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
   150
  by presburger
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   151
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   152
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
   153
by presburger
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   154
21263
wenzelm
parents: 21256
diff changeset
   155
lemma even_nat_equiv_def: "even (x::nat) = (x mod Suc (Suc 0) = 0)"
23522
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
   156
  by presburger
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
lemma odd_nat_equiv_def: "odd (x::nat) = (x mod Suc (Suc 0) = Suc 0)"
23522
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
   159
  by presburger
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   160
21263
wenzelm
parents: 21256
diff changeset
   161
lemma even_nat_div_two_times_two: "even (x::nat) ==>
23522
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
   162
    Suc (Suc 0) * (x div Suc (Suc 0)) = x" by presburger
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   163
21263
wenzelm
parents: 21256
diff changeset
   164
lemma odd_nat_div_two_times_two_plus_one: "odd (x::nat) ==>
23522
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
   165
    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
   166
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   167
lemma even_nat_equiv_def2: "even (x::nat) = (EX y. x = Suc (Suc 0) * y)"
23522
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
   168
  by presburger
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   169
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   170
lemma odd_nat_equiv_def2: "odd (x::nat) = (EX y. x = Suc(Suc (Suc 0) * y))"
23522
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
   171
  by presburger
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   172
25600
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   173
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   174
subsection {* Parity and powers *}
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   175
21263
wenzelm
parents: 21256
diff changeset
   176
lemma  minus_one_even_odd_power:
wenzelm
parents: 21256
diff changeset
   177
     "(even x --> (- 1::'a::{comm_ring_1,recpower})^x = 1) &
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   178
      (odd x --> (- 1::'a)^x = - 1)"
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   179
  apply (induct x)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   180
  apply (rule conjI)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   181
  apply simp
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   182
  apply (insert even_nat_zero, blast)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   183
  apply (simp add: power_Suc)
21263
wenzelm
parents: 21256
diff changeset
   184
  done
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 minus_one_even_power [simp]:
21263
wenzelm
parents: 21256
diff changeset
   187
    "even x ==> (- 1::'a::{comm_ring_1,recpower})^x = 1"
wenzelm
parents: 21256
diff changeset
   188
  using minus_one_even_odd_power by blast
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   189
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   190
lemma minus_one_odd_power [simp]:
21263
wenzelm
parents: 21256
diff changeset
   191
    "odd x ==> (- 1::'a::{comm_ring_1,recpower})^x = - 1"
wenzelm
parents: 21256
diff changeset
   192
  using minus_one_even_odd_power by blast
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   193
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   194
lemma neg_one_even_odd_power:
21263
wenzelm
parents: 21256
diff changeset
   195
     "(even x --> (-1::'a::{number_ring,recpower})^x = 1) &
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   196
      (odd x --> (-1::'a)^x = -1)"
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   197
  apply (induct x)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   198
  apply (simp, simp add: power_Suc)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   199
  done
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   200
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   201
lemma neg_one_even_power [simp]:
21263
wenzelm
parents: 21256
diff changeset
   202
    "even x ==> (-1::'a::{number_ring,recpower})^x = 1"
wenzelm
parents: 21256
diff changeset
   203
  using neg_one_even_odd_power by blast
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   204
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   205
lemma neg_one_odd_power [simp]:
21263
wenzelm
parents: 21256
diff changeset
   206
    "odd x ==> (-1::'a::{number_ring,recpower})^x = -1"
wenzelm
parents: 21256
diff changeset
   207
  using neg_one_even_odd_power by blast
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   208
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   209
lemma neg_power_if:
21263
wenzelm
parents: 21256
diff changeset
   210
     "(-x::'a::{comm_ring_1,recpower}) ^ n =
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   211
      (if even n then (x ^ n) else -(x ^ n))"
21263
wenzelm
parents: 21256
diff changeset
   212
  apply (induct n)
wenzelm
parents: 21256
diff changeset
   213
  apply (simp_all split: split_if_asm add: power_Suc)
wenzelm
parents: 21256
diff changeset
   214
  done
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   215
21263
wenzelm
parents: 21256
diff changeset
   216
lemma zero_le_even_power: "even n ==>
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   217
    0 <= (x::'a::{recpower,ordered_ring_strict}) ^ n"
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   218
  apply (simp add: even_nat_equiv_def2)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   219
  apply (erule exE)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   220
  apply (erule ssubst)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   221
  apply (subst power_add)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   222
  apply (rule zero_le_square)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   223
  done
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   224
21263
wenzelm
parents: 21256
diff changeset
   225
lemma zero_le_odd_power: "odd n ==>
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   226
    (0 <= (x::'a::{recpower,ordered_idom}) ^ n) = (0 <= x)"
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   227
  apply (simp add: odd_nat_equiv_def2)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   228
  apply (erule exE)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   229
  apply (erule ssubst)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   230
  apply (subst power_Suc)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   231
  apply (subst power_add)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   232
  apply (subst zero_le_mult_iff)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   233
  apply auto
25162
ad4d5365d9d8 went back to >0
nipkow
parents: 25134
diff changeset
   234
  apply (subgoal_tac "x = 0 & y > 0")
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   235
  apply (erule conjE, assumption)
21263
wenzelm
parents: 21256
diff changeset
   236
  apply (subst power_eq_0_iff [symmetric])
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   237
  apply (subgoal_tac "0 <= x^y * x^y")
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   238
  apply simp
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   239
  apply (rule zero_le_square)+
21263
wenzelm
parents: 21256
diff changeset
   240
  done
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   241
23522
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
   242
lemma zero_le_power_eq[presburger]: "(0 <= (x::'a::{recpower,ordered_idom}) ^ n) =
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   243
    (even n | (odd n & 0 <= x))"
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   244
  apply auto
21263
wenzelm
parents: 21256
diff changeset
   245
  apply (subst zero_le_odd_power [symmetric])
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   246
  apply assumption+
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   247
  apply (erule zero_le_even_power)
21263
wenzelm
parents: 21256
diff changeset
   248
  apply (subst zero_le_odd_power)
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   249
  apply assumption+
21263
wenzelm
parents: 21256
diff changeset
   250
  done
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   251
23522
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
   252
lemma zero_less_power_eq[presburger]: "(0 < (x::'a::{recpower,ordered_idom}) ^ n) =
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   253
    (n = 0 | (even n & x ~= 0) | (odd n & 0 < x))"
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   254
  apply (rule iffI)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   255
  apply clarsimp
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   256
  apply (rule conjI)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   257
  apply clarsimp
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   258
  apply (rule ccontr)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   259
  apply (subgoal_tac "~ (0 <= x^n)")
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   260
  apply simp
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   261
  apply (subst zero_le_odd_power)
21263
wenzelm
parents: 21256
diff changeset
   262
  apply assumption
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   263
  apply simp
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   264
  apply (rule notI)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   265
  apply (simp add: power_0_left)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   266
  apply (rule notI)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   267
  apply (simp add: power_0_left)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   268
  apply auto
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   269
  apply (subgoal_tac "0 <= x^n")
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   270
  apply (frule order_le_imp_less_or_eq)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   271
  apply simp
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   272
  apply (erule zero_le_even_power)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   273
  apply (subgoal_tac "0 <= x^n")
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   274
  apply (frule order_le_imp_less_or_eq)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   275
  apply auto
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   276
  apply (subst zero_le_odd_power)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   277
  apply assumption
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   278
  apply (erule order_less_imp_le)
21263
wenzelm
parents: 21256
diff changeset
   279
  done
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   280
23522
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
   281
lemma power_less_zero_eq[presburger]: "((x::'a::{recpower,ordered_idom}) ^ n < 0) =
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
   282
    (odd n & x < 0)" 
21263
wenzelm
parents: 21256
diff changeset
   283
  apply (subst linorder_not_le [symmetric])+
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   284
  apply (subst zero_le_power_eq)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   285
  apply auto
21263
wenzelm
parents: 21256
diff changeset
   286
  done
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   287
23522
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
   288
lemma power_le_zero_eq[presburger]: "((x::'a::{recpower,ordered_idom}) ^ n <= 0) =
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   289
    (n ~= 0 & ((odd n & x <= 0) | (even n & x = 0)))"
21263
wenzelm
parents: 21256
diff changeset
   290
  apply (subst linorder_not_less [symmetric])+
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   291
  apply (subst zero_less_power_eq)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   292
  apply auto
21263
wenzelm
parents: 21256
diff changeset
   293
  done
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   294
21263
wenzelm
parents: 21256
diff changeset
   295
lemma power_even_abs: "even n ==>
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   296
    (abs (x::'a::{recpower,ordered_idom}))^n = x^n"
21263
wenzelm
parents: 21256
diff changeset
   297
  apply (subst power_abs [symmetric])
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   298
  apply (simp add: zero_le_even_power)
21263
wenzelm
parents: 21256
diff changeset
   299
  done
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   300
23522
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
   301
lemma zero_less_power_nat_eq[presburger]: "(0 < (x::nat) ^ n) = (n = 0 | 0 < x)"
21263
wenzelm
parents: 21256
diff changeset
   302
  by (induct n) auto
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   303
21263
wenzelm
parents: 21256
diff changeset
   304
lemma power_minus_even [simp]: "even n ==>
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   305
    (- x)^n = (x^n::'a::{recpower,comm_ring_1})"
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   306
  apply (subst power_minus)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   307
  apply simp
21263
wenzelm
parents: 21256
diff changeset
   308
  done
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   309
21263
wenzelm
parents: 21256
diff changeset
   310
lemma power_minus_odd [simp]: "odd n ==>
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   311
    (- x)^n = - (x^n::'a::{recpower,comm_ring_1})"
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   312
  apply (subst power_minus)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   313
  apply simp
21263
wenzelm
parents: 21256
diff changeset
   314
  done
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   315
21263
wenzelm
parents: 21256
diff changeset
   316
25600
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   317
subsection {* General Lemmas About Division *}
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   318
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   319
lemma Suc_times_mod_eq: "1<k ==> Suc (k * m) mod k = 1" 
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   320
apply (induct "m")
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   321
apply (simp_all add: mod_Suc)
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   322
done
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   323
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   324
declare Suc_times_mod_eq [of "number_of w", standard, simp]
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   325
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   326
lemma [simp]: "n div k \<le> (Suc n) div k"
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   327
by (simp add: div_le_mono) 
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   328
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   329
lemma Suc_n_div_2_gt_zero [simp]: "(0::nat) < n ==> 0 < (n + 1) div 2"
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   330
by arith
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   331
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   332
lemma div_2_gt_zero [simp]: "(1::nat) < n ==> 0 < n div 2" 
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   333
by arith
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   334
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   335
lemma mod_mult_self3 [simp]: "(k*n + m) mod n = m mod (n::nat)"
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   336
by (simp add: mult_ac add_ac)
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   337
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   338
lemma mod_mult_self4 [simp]: "Suc (k*n + m) mod n = Suc m mod n"
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   339
proof -
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   340
  have "Suc (k * n + m) mod n = (k * n + Suc m) mod n" by simp
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   341
  also have "... = Suc m mod n" by (rule mod_mult_self3) 
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   342
  finally show ?thesis .
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   343
qed
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   344
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   345
lemma mod_Suc_eq_Suc_mod: "Suc m mod n = Suc (m mod n) mod n"
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   346
apply (subst mod_Suc [of m]) 
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   347
apply (subst mod_Suc [of "m mod n"], simp) 
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   348
done
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   349
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   350
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   351
subsection {* More Even/Odd Results *}
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   352
 
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   353
lemma even_mult_two_ex: "even(n) = (\<exists>m::nat. n = 2*m)"
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   354
by (simp add: even_nat_equiv_def2 numeral_2_eq_2)
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 odd_Suc_mult_two_ex: "odd(n) = (\<exists>m. n = Suc (2*m))"
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   357
by (simp add: odd_nat_equiv_def2 numeral_2_eq_2)
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   358
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   359
lemma even_add [simp]: "even(m + n::nat) = (even m = even n)" 
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   360
by auto
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   361
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   362
lemma odd_add [simp]: "odd(m + n::nat) = (odd m \<noteq> odd n)"
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   363
by auto
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   364
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   365
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
   366
    (a mod c + Suc 0 mod c) div c" 
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   367
  apply (subgoal_tac "Suc a = a + Suc 0")
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   368
  apply (erule ssubst)
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   369
  apply (rule div_add1_eq, simp)
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   370
  done
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   371
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   372
lemma lemma_even_div2 [simp]: "even (n::nat) ==> (n + 1) div 2 = n div 2"
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   373
apply (simp add: numeral_2_eq_2) 
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   374
apply (subst div_Suc)  
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   375
apply (simp add: even_nat_mod_two_eq_zero) 
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   376
done
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   377
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   378
lemma lemma_not_even_div2 [simp]: "~even n ==> (n + 1) div 2 = Suc (n div 2)"
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   379
apply (simp add: numeral_2_eq_2) 
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   380
apply (subst div_Suc)  
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   381
apply (simp add: odd_nat_mod_two_eq_one) 
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   382
done
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   383
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   384
lemma even_num_iff: "0 < n ==> even n = (~ even(n - 1 :: nat))" 
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   385
by (case_tac "n", auto)
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   386
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   387
lemma even_even_mod_4_iff: "even (n::nat) = even (n mod 4)"
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   388
apply (induct n, simp)
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   389
apply (subst mod_Suc, simp) 
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   390
done
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   391
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   392
lemma lemma_odd_mod_4_div_2: "n mod 4 = (3::nat) ==> odd((n - 1) div 2)"
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   393
apply (rule_tac t = n and n1 = 4 in mod_div_equality [THEN subst])
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   394
apply (simp add: even_num_iff)
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   395
done
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   396
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   397
lemma lemma_even_mod_4_div_2: "n mod 4 = (1::nat) ==> even ((n - 1) div 2)"
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   398
by (rule_tac t = n and n1 = 4 in mod_div_equality [THEN subst], simp)
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   399
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   400
21263
wenzelm
parents: 21256
diff changeset
   401
text {* Simplify, when the exponent is a numeral *}
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   402
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   403
lemmas power_0_left_number_of = power_0_left [of "number_of w", standard]
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   404
declare power_0_left_number_of [simp]
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   405
21263
wenzelm
parents: 21256
diff changeset
   406
lemmas zero_le_power_eq_number_of [simp] =
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   407
    zero_le_power_eq [of _ "number_of w", standard]
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   408
21263
wenzelm
parents: 21256
diff changeset
   409
lemmas zero_less_power_eq_number_of [simp] =
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   410
    zero_less_power_eq [of _ "number_of w", standard]
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   411
21263
wenzelm
parents: 21256
diff changeset
   412
lemmas power_le_zero_eq_number_of [simp] =
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   413
    power_le_zero_eq [of _ "number_of w", standard]
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   414
21263
wenzelm
parents: 21256
diff changeset
   415
lemmas power_less_zero_eq_number_of [simp] =
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   416
    power_less_zero_eq [of _ "number_of w", standard]
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   417
21263
wenzelm
parents: 21256
diff changeset
   418
lemmas zero_less_power_nat_eq_number_of [simp] =
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   419
    zero_less_power_nat_eq [of _ "number_of w", standard]
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   420
21263
wenzelm
parents: 21256
diff changeset
   421
lemmas power_eq_0_iff_number_of [simp] = power_eq_0_iff [of _ "number_of w", standard]
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   422
21263
wenzelm
parents: 21256
diff changeset
   423
lemmas power_even_abs_number_of [simp] = power_even_abs [of "number_of w" _, standard]
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   424
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   425
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   426
subsection {* An Equivalence for @{term [source] "0 \<le> a^n"} *}
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   427
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   428
lemma even_power_le_0_imp_0:
21263
wenzelm
parents: 21256
diff changeset
   429
    "a ^ (2*k) \<le> (0::'a::{ordered_idom,recpower}) ==> a=0"
wenzelm
parents: 21256
diff changeset
   430
  by (induct k) (auto simp add: zero_le_mult_iff mult_le_0_iff power_Suc)
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   431
23522
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
   432
lemma zero_le_power_iff[presburger]:
21263
wenzelm
parents: 21256
diff changeset
   433
  "(0 \<le> a^n) = (0 \<le> (a::'a::{ordered_idom,recpower}) | even n)"
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   434
proof cases
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   435
  assume even: "even n"
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   436
  then obtain k where "n = 2*k"
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   437
    by (auto simp add: even_nat_equiv_def2 numeral_2_eq_2)
21263
wenzelm
parents: 21256
diff changeset
   438
  thus ?thesis by (simp add: zero_le_even_power even)
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   439
next
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   440
  assume odd: "odd n"
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   441
  then obtain k where "n = Suc(2*k)"
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   442
    by (auto simp add: odd_nat_equiv_def2 numeral_2_eq_2)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   443
  thus ?thesis
21263
wenzelm
parents: 21256
diff changeset
   444
    by (auto simp add: power_Suc zero_le_mult_iff zero_le_even_power
wenzelm
parents: 21256
diff changeset
   445
             dest!: even_power_le_0_imp_0)
wenzelm
parents: 21256
diff changeset
   446
qed
wenzelm
parents: 21256
diff changeset
   447
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   448
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   449
subsection {* Miscellaneous *}
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   450
25600
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   451
lemma odd_pos: "odd (n::nat) \<Longrightarrow> 0 < n"
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   452
  by (cases n, simp_all)
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   453
23522
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
   454
lemma [presburger]:"(x + 1) div 2 = x div 2 \<longleftrightarrow> even (x::int)" by presburger
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
   455
lemma [presburger]: "(x + 1) div 2 = x div 2 + 1 \<longleftrightarrow> odd (x::int)" by presburger
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
   456
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
   457
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
   458
23522
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
   459
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
   460
lemma [presburger]: "(Suc x) div Suc (Suc 0) = x div Suc (Suc 0) \<longleftrightarrow> even x" by presburger
21263
wenzelm
parents: 21256
diff changeset
   461
lemma even_nat_plus_one_div_two: "even (x::nat) ==>
23522
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
   462
    (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
   463
21263
wenzelm
parents: 21256
diff changeset
   464
lemma odd_nat_plus_one_div_two: "odd (x::nat) ==>
23522
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
   465
    (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
   466
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   467
end