src/HOL/Parity.thy
author haftmann
Thu, 16 Oct 2014 19:26:14 +0200
changeset 58688 ddd124805260
parent 58687 5469874b0228
child 58689 ee5bf401cfa7
permissions -rw-r--r--
restructured
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
58678
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    12
subsection {* Preliminaries about divisibility on @{typ nat} and @{typ int} *}
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    13
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    14
lemma two_dvd_Suc_Suc_iff [simp]:
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    15
  "2 dvd Suc (Suc n) \<longleftrightarrow> 2 dvd n"
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    16
  using dvd_add_triv_right_iff [of 2 n] by simp
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    17
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    18
lemma two_dvd_Suc_iff:
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    19
  "2 dvd Suc n \<longleftrightarrow> \<not> 2 dvd n"
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    20
  by (induct n) auto
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    21
58687
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
    22
lemma two_dvd_diff_nat_iff:
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
    23
  fixes m n :: nat
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
    24
  shows "2 dvd m - n \<longleftrightarrow> m < n \<or> 2 dvd m + n"
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
    25
proof (cases "n \<le> m")
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
    26
  case True
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
    27
  then have "m - n + n * 2 = m + n" by simp
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
    28
  moreover have "2 dvd m - n \<longleftrightarrow> 2 dvd m - n + n * 2" by simp
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
    29
  ultimately have "2 dvd m - n \<longleftrightarrow> 2 dvd m + n" by (simp only:)
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
    30
  then show ?thesis by auto
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
    31
next
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
    32
  case False
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
    33
  then show ?thesis by simp
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
    34
qed 
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
    35
  
58678
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    36
lemma two_dvd_diff_iff:
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    37
  fixes k l :: int
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    38
  shows "2 dvd k - l \<longleftrightarrow> 2 dvd k + l"
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    39
  using dvd_add_times_triv_right_iff [of 2 "k - l" l] by (simp add: ac_simps)
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    40
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    41
lemma two_dvd_abs_add_iff:
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    42
  fixes k l :: int
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    43
  shows "2 dvd \<bar>k\<bar> + l \<longleftrightarrow> 2 dvd k + l"
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    44
  by (cases "k \<ge> 0") (simp_all add: two_dvd_diff_iff ac_simps)
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    45
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    46
lemma two_dvd_add_abs_iff:
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    47
  fixes k l :: int
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    48
  shows "2 dvd k + \<bar>l\<bar> \<longleftrightarrow> 2 dvd k + l"
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    49
  using two_dvd_abs_add_iff [of l k] by (simp add: ac_simps)
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    50
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    51
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    52
subsection {* Ring structures with parity *}
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    53
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    54
class semiring_parity = semiring_dvd + semiring_numeral +
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    55
  assumes two_not_dvd_one [simp]: "\<not> 2 dvd 1"
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    56
  assumes not_dvd_not_dvd_dvd_add: "\<not> 2 dvd a \<Longrightarrow> \<not> 2 dvd b \<Longrightarrow> 2 dvd a + b"
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    57
  assumes two_is_prime: "2 dvd a * b \<Longrightarrow> 2 dvd a \<or> 2 dvd b"
58680
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
    58
  assumes not_dvd_ex_decrement: "\<not> 2 dvd a \<Longrightarrow> \<exists>b. a = b + 1"
58678
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    59
begin
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    60
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    61
lemma two_dvd_plus_one_iff [simp]:
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    62
  "2 dvd a + 1 \<longleftrightarrow> \<not> 2 dvd a"
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    63
  by (auto simp add: dvd_add_right_iff intro: not_dvd_not_dvd_dvd_add)
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    64
58680
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
    65
lemma not_two_dvdE [elim?]:
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
    66
  assumes "\<not> 2 dvd a"
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
    67
  obtains b where "a = 2 * b + 1"
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
    68
proof -
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
    69
  from assms obtain b where *: "a = b + 1"
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
    70
    by (blast dest: not_dvd_ex_decrement)
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
    71
  with assms have "2 dvd b + 2" by simp
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
    72
  then have "2 dvd b" by simp
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
    73
  then obtain c where "b = 2 * c" ..
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
    74
  with * have "a = 2 * c + 1" by simp
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
    75
  with that show thesis .
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
    76
qed
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
    77
58678
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    78
end
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    79
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    80
instance nat :: semiring_parity
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    81
proof
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    82
  show "\<not> (2 :: nat) dvd 1"
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    83
    by (rule notI, erule dvdE) simp
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    84
next
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    85
  fix m n :: nat
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    86
  assume "\<not> 2 dvd m"
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    87
  moreover assume "\<not> 2 dvd n"
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    88
  ultimately have *: "2 dvd Suc m \<and> 2 dvd Suc n"
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    89
    by (simp add: two_dvd_Suc_iff)
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    90
  then have "2 dvd Suc m + Suc n"
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    91
    by (blast intro: dvd_add)
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    92
  also have "Suc m + Suc n = m + n + 2"
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    93
    by simp
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    94
  finally show "2 dvd m + n"
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    95
    using dvd_add_triv_right_iff [of 2 "m + n"] by simp
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    96
next
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    97
  fix m n :: nat
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    98
  assume *: "2 dvd m * n"
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    99
  show "2 dvd m \<or> 2 dvd n"
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   100
  proof (rule disjCI)
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   101
    assume "\<not> 2 dvd n"
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   102
    then have "2 dvd Suc n" by (simp add: two_dvd_Suc_iff)
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   103
    then obtain r where "Suc n = 2 * r" ..
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   104
    moreover from * obtain s where "m * n = 2 * s" ..
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   105
    then have "2 * s + m = m * Suc n" by simp
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   106
    ultimately have " 2 * s + m = 2 * (m * r)" by simp
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   107
    then have "m = 2 * (m * r - s)" by simp
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   108
    then show "2 dvd m" ..
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   109
  qed
58680
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   110
next
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   111
  fix n :: nat
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   112
  assume "\<not> 2 dvd n"
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   113
  then show "\<exists>m. n = m + 1"
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   114
    by (cases n) simp_all
58678
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   115
qed
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   116
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   117
class ring_parity = comm_ring_1 + semiring_parity
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   118
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   119
instance int :: ring_parity
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   120
proof
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   121
  show "\<not> (2 :: int) dvd 1" by (simp add: dvd_int_unfold_dvd_nat)
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   122
  fix k l :: int
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   123
  assume "\<not> 2 dvd k"
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   124
  moreover assume "\<not> 2 dvd l"
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   125
  ultimately have "2 dvd nat \<bar>k\<bar> + nat \<bar>l\<bar>" 
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   126
    by (auto simp add: dvd_int_unfold_dvd_nat intro: not_dvd_not_dvd_dvd_add)
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   127
  then have "2 dvd \<bar>k\<bar> + \<bar>l\<bar>"
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   128
    by (simp add: dvd_int_unfold_dvd_nat nat_add_distrib)
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   129
  then show "2 dvd k + l"
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   130
    by (simp add: two_dvd_abs_add_iff two_dvd_add_abs_iff)
58680
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   131
next
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   132
  fix k l :: int
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   133
  assume "2 dvd k * l"
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   134
  then show "2 dvd k \<or> 2 dvd l"
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   135
    by (simp add: dvd_int_unfold_dvd_nat two_is_prime nat_abs_mult_distrib)
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   136
next
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   137
  fix k :: int
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   138
  have "k = (k - 1) + 1" by simp
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   139
  then show "\<exists>l. k = l + 1" ..
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   140
qed
58678
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   141
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   142
context semiring_div_parity
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   143
begin
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   144
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   145
subclass semiring_parity
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   146
proof (unfold_locales, unfold dvd_eq_mod_eq_0 not_mod_2_eq_0_eq_1)
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   147
  fix a b c
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   148
  show "(c * a + b) mod a = 0 \<longleftrightarrow> b mod a = 0"
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   149
    by simp
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   150
next
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   151
  fix a b c
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   152
  assume "(b + c) mod a = 0"
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   153
  with mod_add_eq [of b c a]
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   154
  have "(b mod a + c mod a) mod a = 0"
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   155
    by simp
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   156
  moreover assume "b mod a = 0"
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   157
  ultimately show "c mod a = 0"
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   158
    by simp
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   159
next
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   160
  show "1 mod 2 = 1"
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   161
    by (fact one_mod_two_eq_one)
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   162
next
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   163
  fix a b
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   164
  assume "a mod 2 = 1"
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   165
  moreover assume "b mod 2 = 1"
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   166
  ultimately show "(a + b) mod 2 = 0"
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   167
    using mod_add_eq [of a b 2] by simp
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   168
next
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   169
  fix a b
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   170
  assume "(a * b) mod 2 = 0"
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   171
  then have "(a mod 2) * (b mod 2) = 0"
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   172
    by (cases "a mod 2 = 0") (simp_all add: mod_mult_eq [of a b 2])
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   173
  then show "a mod 2 = 0 \<or> b mod 2 = 0"
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   174
    by (rule divisors_zero)
58680
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   175
next
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   176
  fix a
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   177
  assume "a mod 2 = 1"
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   178
  then have "a = a div 2 * 2 + 1" using mod_div_equality [of a 2] by simp
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   179
  then show "\<exists>b. a = b + 1" ..
58678
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   180
qed
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   181
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   182
end
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   183
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   184
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   185
subsection {* Dedicated @{text even}/@{text odd} predicate *}
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   186
58680
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   187
subsubsection {* Properties *}
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   188
58678
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   189
context semiring_parity
54227
63b441f49645 moving generic lemmas out of theory parity, disregarding some unused auxiliary lemmas;
haftmann
parents: 47225
diff changeset
   190
begin
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   191
54228
229282d53781 purely algebraic foundation for even/odd
haftmann
parents: 54227
diff changeset
   192
definition even :: "'a \<Rightarrow> bool"
229282d53781 purely algebraic foundation for even/odd
haftmann
parents: 54227
diff changeset
   193
where
58645
94bef115c08f more foundational definition for predicate even
haftmann
parents: 54489
diff changeset
   194
  [algebra]: "even a \<longleftrightarrow> 2 dvd a"
54228
229282d53781 purely algebraic foundation for even/odd
haftmann
parents: 54227
diff changeset
   195
58678
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   196
abbreviation odd :: "'a \<Rightarrow> bool"
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   197
where
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   198
  "odd a \<equiv> \<not> even a"
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   199
58681
a478a0742a8e legacy cleanup
haftmann
parents: 58680
diff changeset
   200
lemma oddE [elim?]:
58680
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   201
  assumes "odd a"
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   202
  obtains b where "a = 2 * b + 1"
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   203
proof -
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   204
  from assms have "\<not> 2 dvd a" by (simp add: even_def)
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   205
  then show thesis using that by (rule not_two_dvdE)
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   206
qed
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   207
  
58678
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   208
lemma even_times_iff [simp, presburger, algebra]:
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   209
  "even (a * b) \<longleftrightarrow> even a \<or> even b"
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   210
  by (auto simp add: even_def dest: two_is_prime)
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   211
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   212
lemma even_zero [simp]:
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   213
  "even 0"
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   214
  by (simp add: even_def)
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   215
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   216
lemma odd_one [simp]:
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   217
  "odd 1"
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   218
  by (simp add: even_def)
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   219
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   220
lemma even_numeral [simp]:
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   221
  "even (numeral (Num.Bit0 n))"
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   222
proof -
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   223
  have "even (2 * numeral n)"
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   224
    unfolding even_times_iff by (simp add: even_def)
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   225
  then have "even (numeral n + numeral n)"
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   226
    unfolding mult_2 .
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   227
  then show ?thesis
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   228
    unfolding numeral.simps .
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   229
qed
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   230
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   231
lemma odd_numeral [simp]:
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   232
  "odd (numeral (Num.Bit1 n))"
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   233
proof
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   234
  assume "even (numeral (num.Bit1 n))"
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   235
  then have "even (numeral n + numeral n + 1)"
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   236
    unfolding numeral.simps .
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   237
  then have "even (2 * numeral n + 1)"
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   238
    unfolding mult_2 .
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   239
  then have "2 dvd numeral n * 2 + 1"
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   240
    unfolding even_def by (simp add: ac_simps)
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   241
  with dvd_add_times_triv_left_iff [of 2 "numeral n" 1]
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   242
    have "2 dvd 1"
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   243
    by simp
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   244
  then show False by simp
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   245
qed
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   246
58680
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   247
lemma even_add [simp]:
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   248
  "even (a + b) \<longleftrightarrow> (even a \<longleftrightarrow> even b)"
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   249
  by (auto simp add: even_def dvd_add_right_iff dvd_add_left_iff not_dvd_not_dvd_dvd_add)
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   250
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   251
lemma odd_add [simp]:
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   252
  "odd (a + b) \<longleftrightarrow> (\<not> (odd a \<longleftrightarrow> odd b))"
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   253
  by simp
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   254
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   255
lemma even_power [simp, presburger]:
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   256
  "even (a ^ n) \<longleftrightarrow> even a \<and> n \<noteq> 0"
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   257
  by (induct n) auto
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   258
58678
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   259
end
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   260
58679
33c90658448a more algebraic deductions for facts on even/odd
haftmann
parents: 58678
diff changeset
   261
context ring_parity
33c90658448a more algebraic deductions for facts on even/odd
haftmann
parents: 58678
diff changeset
   262
begin
33c90658448a more algebraic deductions for facts on even/odd
haftmann
parents: 58678
diff changeset
   263
33c90658448a more algebraic deductions for facts on even/odd
haftmann
parents: 58678
diff changeset
   264
lemma even_minus [simp, presburger, algebra]:
33c90658448a more algebraic deductions for facts on even/odd
haftmann
parents: 58678
diff changeset
   265
  "even (- a) \<longleftrightarrow> even a"
33c90658448a more algebraic deductions for facts on even/odd
haftmann
parents: 58678
diff changeset
   266
  by (simp add: even_def)
33c90658448a more algebraic deductions for facts on even/odd
haftmann
parents: 58678
diff changeset
   267
58680
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   268
lemma even_diff [simp]:
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   269
  "even (a - b) \<longleftrightarrow> even (a + b)"
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   270
  using even_add [of a "- b"] by simp
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   271
58679
33c90658448a more algebraic deductions for facts on even/odd
haftmann
parents: 58678
diff changeset
   272
end
33c90658448a more algebraic deductions for facts on even/odd
haftmann
parents: 58678
diff changeset
   273
58678
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   274
context semiring_div_parity
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   275
begin
58645
94bef115c08f more foundational definition for predicate even
haftmann
parents: 54489
diff changeset
   276
94bef115c08f more foundational definition for predicate even
haftmann
parents: 54489
diff changeset
   277
lemma even_iff_mod_2_eq_zero [presburger]:
94bef115c08f more foundational definition for predicate even
haftmann
parents: 54489
diff changeset
   278
  "even a \<longleftrightarrow> a mod 2 = 0"
54228
229282d53781 purely algebraic foundation for even/odd
haftmann
parents: 54227
diff changeset
   279
  by (simp add: even_def dvd_eq_mod_eq_0)
229282d53781 purely algebraic foundation for even/odd
haftmann
parents: 54227
diff changeset
   280
54227
63b441f49645 moving generic lemmas out of theory parity, disregarding some unused auxiliary lemmas;
haftmann
parents: 47225
diff changeset
   281
end
63b441f49645 moving generic lemmas out of theory parity, disregarding some unused auxiliary lemmas;
haftmann
parents: 47225
diff changeset
   282
58680
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   283
58687
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   284
subsubsection {* Particularities for @{typ nat} and @{typ int} *}
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   285
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   286
lemma even_Suc [simp, presburger, algebra]:
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   287
  "even (Suc n) = odd n"
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   288
  by (simp add: even_def two_dvd_Suc_iff)
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   289
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   290
lemma even_diff_nat [simp]:
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   291
  fixes m n :: nat
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   292
  shows "even (m - n) \<longleftrightarrow> m < n \<or> even (m + n)"
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   293
  by (simp add: even_def two_dvd_diff_nat_iff)
58680
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   294
58679
33c90658448a more algebraic deductions for facts on even/odd
haftmann
parents: 58678
diff changeset
   295
lemma even_int_iff:
33c90658448a more algebraic deductions for facts on even/odd
haftmann
parents: 58678
diff changeset
   296
  "even (int n) \<longleftrightarrow> even n"
33c90658448a more algebraic deductions for facts on even/odd
haftmann
parents: 58678
diff changeset
   297
  by (simp add: even_def dvd_int_iff)
33318
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 31718
diff changeset
   298
58687
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   299
lemma even_nat_iff:
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   300
  "0 \<le> k \<Longrightarrow> even (nat k) \<longleftrightarrow> even k"
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   301
  by (simp add: even_int_iff [symmetric])
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   302
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   303
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   304
subsubsection {* Tools setup *}
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   305
58679
33c90658448a more algebraic deductions for facts on even/odd
haftmann
parents: 58678
diff changeset
   306
declare transfer_morphism_int_nat [transfer add return:
33c90658448a more algebraic deductions for facts on even/odd
haftmann
parents: 58678
diff changeset
   307
  even_int_iff
33318
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 31718
diff changeset
   308
]
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   309
58679
33c90658448a more algebraic deductions for facts on even/odd
haftmann
parents: 58678
diff changeset
   310
lemma [presburger]:
33c90658448a more algebraic deductions for facts on even/odd
haftmann
parents: 58678
diff changeset
   311
  "even n \<longleftrightarrow> even (int n)"
33c90658448a more algebraic deductions for facts on even/odd
haftmann
parents: 58678
diff changeset
   312
  using even_int_iff [of n] by simp
25600
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   313
58687
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   314
lemma (in semiring_parity) [presburger]:
58680
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   315
  "even (a + b) \<longleftrightarrow> even a \<and> even b \<or> odd a \<and> odd b"
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   316
  by auto
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   317
58687
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   318
lemma [presburger, algebra]:
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   319
  fixes m n :: nat
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   320
  shows "even (m - n) \<longleftrightarrow> m < n \<or> even m \<and> even n \<or> odd m \<and> odd n"
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   321
  by auto
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   322
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   323
lemma [presburger, algebra]:
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   324
  fixes m n :: nat
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   325
  shows "even (m ^ n) \<longleftrightarrow> even m \<and> 0 < n"
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   326
  by simp
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   327
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   328
lemma [presburger]:
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   329
  fixes k :: int
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   330
  shows "(k + 1) div 2 = k div 2 \<longleftrightarrow> even k"
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   331
  by presburger
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   332
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   333
lemma [presburger]:
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   334
  fixes k :: int
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   335
  shows "(k + 1) div 2 = k div 2 + 1 \<longleftrightarrow> odd k"
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   336
  by presburger
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   337
  
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   338
lemma [presburger]:
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   339
  "Suc n div Suc (Suc 0) = n div Suc (Suc 0) \<longleftrightarrow> even n"
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   340
  by presburger
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   341
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   342
58688
ddd124805260 restructured
haftmann
parents: 58687
diff changeset
   343
subsubsection {* Miscellaneous *}
58680
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   344
58688
ddd124805260 restructured
haftmann
parents: 58687
diff changeset
   345
lemma even_nat_plus_one_div_two:
ddd124805260 restructured
haftmann
parents: 58687
diff changeset
   346
  "even (x::nat) ==> (Suc x) div Suc (Suc 0) = x div Suc (Suc 0)"
ddd124805260 restructured
haftmann
parents: 58687
diff changeset
   347
  by presburger
58680
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   348
58688
ddd124805260 restructured
haftmann
parents: 58687
diff changeset
   349
lemma odd_nat_plus_one_div_two:
ddd124805260 restructured
haftmann
parents: 58687
diff changeset
   350
  "odd (x::nat) ==> (Suc x) div Suc (Suc 0) = Suc (x div Suc (Suc 0))"
ddd124805260 restructured
haftmann
parents: 58687
diff changeset
   351
  by presburger
58680
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   352
58687
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   353
lemma even_nat_mod_two_eq_zero:
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   354
  "even (x::nat) ==> x mod Suc (Suc 0) = 0"
58680
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   355
  by presburger
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   356
58687
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   357
lemma odd_nat_mod_two_eq_one:
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   358
  "odd (x::nat) ==> x mod Suc (Suc 0) = Suc 0"
58680
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   359
  by presburger
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   360
58687
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   361
lemma even_nat_equiv_def:
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   362
  "even (x::nat) = (x mod Suc (Suc 0) = 0)"
58680
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   363
  by presburger
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   364
58687
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   365
lemma odd_nat_equiv_def:
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   366
  "odd (x::nat) = (x mod Suc (Suc 0) = Suc 0)"
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   367
  by presburger
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   368
58687
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   369
lemma even_nat_div_two_times_two:
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   370
  "even (x::nat) ==> Suc (Suc 0) * (x div Suc (Suc 0)) = x"
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   371
  by presburger
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   372
58687
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   373
lemma odd_nat_div_two_times_two_plus_one:
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   374
  "odd (x::nat) ==> Suc (Suc (Suc 0) * (x div Suc (Suc 0))) = x"
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   375
  by presburger
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   376
58688
ddd124805260 restructured
haftmann
parents: 58687
diff changeset
   377
lemma even_mult_two_ex: "even(n) = (\<exists>m::nat. n = 2*m)" by presburger
ddd124805260 restructured
haftmann
parents: 58687
diff changeset
   378
lemma odd_Suc_mult_two_ex: "odd(n) = (\<exists>m. n = Suc (2*m))" by presburger
ddd124805260 restructured
haftmann
parents: 58687
diff changeset
   379
ddd124805260 restructured
haftmann
parents: 58687
diff changeset
   380
lemma lemma_even_div2 [simp]: "even (n::nat) ==> (n + 1) div 2 = n div 2"
ddd124805260 restructured
haftmann
parents: 58687
diff changeset
   381
  by presburger
ddd124805260 restructured
haftmann
parents: 58687
diff changeset
   382
ddd124805260 restructured
haftmann
parents: 58687
diff changeset
   383
lemma lemma_odd_div2 [simp]: "odd n ==> (n + 1) div 2 = Suc (n div 2)"
ddd124805260 restructured
haftmann
parents: 58687
diff changeset
   384
  by presburger
ddd124805260 restructured
haftmann
parents: 58687
diff changeset
   385
ddd124805260 restructured
haftmann
parents: 58687
diff changeset
   386
lemma even_num_iff: "0 < n ==> even n = (odd (n - 1 :: nat))" by presburger
ddd124805260 restructured
haftmann
parents: 58687
diff changeset
   387
lemma even_even_mod_4_iff: "even (n::nat) = even (n mod 4)" by presburger
ddd124805260 restructured
haftmann
parents: 58687
diff changeset
   388
ddd124805260 restructured
haftmann
parents: 58687
diff changeset
   389
lemma lemma_odd_mod_4_div_2: "n mod 4 = (3::nat) ==> odd((n - 1) div 2)" by presburger
ddd124805260 restructured
haftmann
parents: 58687
diff changeset
   390
ddd124805260 restructured
haftmann
parents: 58687
diff changeset
   391
lemma lemma_even_mod_4_div_2: "n mod 4 = (1::nat) ==> even ((n - 1) div 2)"
ddd124805260 restructured
haftmann
parents: 58687
diff changeset
   392
  by presburger
ddd124805260 restructured
haftmann
parents: 58687
diff changeset
   393
25600
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   394
58680
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   395
subsubsection {* Parity and powers *}
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   396
54228
229282d53781 purely algebraic foundation for even/odd
haftmann
parents: 54227
diff changeset
   397
lemma (in comm_ring_1) neg_power_if:
229282d53781 purely algebraic foundation for even/odd
haftmann
parents: 54227
diff changeset
   398
  "(- a) ^ n = (if even n then (a ^ n) else - (a ^ n))"
229282d53781 purely algebraic foundation for even/odd
haftmann
parents: 54227
diff changeset
   399
  by (induct n) simp_all
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   400
54228
229282d53781 purely algebraic foundation for even/odd
haftmann
parents: 54227
diff changeset
   401
lemma (in comm_ring_1)
54489
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54228
diff changeset
   402
  shows neg_one_even_power [simp]: "even n \<Longrightarrow> (- 1) ^ n = 1"
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54228
diff changeset
   403
  and neg_one_odd_power [simp]: "odd n \<Longrightarrow> (- 1) ^ n = - 1"
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54228
diff changeset
   404
  by (simp_all add: neg_power_if)
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   405
21263
wenzelm
parents: 21256
diff changeset
   406
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
   407
    0 <= (x::'a::{linordered_ring,monoid_mult}) ^ n"
58680
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   408
  apply (simp add: even_def)
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   409
  apply (erule dvdE)
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   410
  apply (erule ssubst)
58680
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   411
  unfolding mult.commute [of 2]
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   412
  unfolding power_mult power2_eq_square
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   413
  apply (rule zero_le_square)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   414
  done
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   415
58681
a478a0742a8e legacy cleanup
haftmann
parents: 58680
diff changeset
   416
lemma zero_le_odd_power:
a478a0742a8e legacy cleanup
haftmann
parents: 58680
diff changeset
   417
  "odd n \<Longrightarrow> 0 \<le> (x::'a::{linordered_idom}) ^ n \<longleftrightarrow> 0 \<le> x"
a478a0742a8e legacy cleanup
haftmann
parents: 58680
diff changeset
   418
  by (erule oddE) (auto simp: power_add zero_le_mult_iff mult_2 order_antisym_conv)
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   419
54227
63b441f49645 moving generic lemmas out of theory parity, disregarding some unused auxiliary lemmas;
haftmann
parents: 47225
diff changeset
   420
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
   421
    (even n | (odd n & 0 <= x))"
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   422
  apply auto
21263
wenzelm
parents: 21256
diff changeset
   423
  apply (subst zero_le_odd_power [symmetric])
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   424
  apply assumption+
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   425
  apply (erule zero_le_even_power)
21263
wenzelm
parents: 21256
diff changeset
   426
  done
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   427
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 33358
diff changeset
   428
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
   429
    (n = 0 | (even n & x ~= 0) | (odd n & 0 < x))"
27668
6eb20b2cecf8 Tuned and simplified proofs
chaieb
parents: 27651
diff changeset
   430
  unfolding order_less_le zero_le_power_eq by auto
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   431
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 33358
diff changeset
   432
lemma power_less_zero_eq[presburger]: "((x::'a::{linordered_idom}) ^ n < 0) =
27668
6eb20b2cecf8 Tuned and simplified proofs
chaieb
parents: 27651
diff changeset
   433
    (odd n & x < 0)"
21263
wenzelm
parents: 21256
diff changeset
   434
  apply (subst linorder_not_le [symmetric])+
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   435
  apply (subst zero_le_power_eq)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   436
  apply auto
21263
wenzelm
parents: 21256
diff changeset
   437
  done
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   438
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 33358
diff changeset
   439
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
   440
    (n ~= 0 & ((odd n & x <= 0) | (even n & x = 0)))"
21263
wenzelm
parents: 21256
diff changeset
   441
  apply (subst linorder_not_less [symmetric])+
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   442
  apply (subst zero_less_power_eq)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   443
  apply auto
21263
wenzelm
parents: 21256
diff changeset
   444
  done
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   445
21263
wenzelm
parents: 21256
diff changeset
   446
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
   447
    (abs (x::'a::{linordered_idom}))^n = x^n"
21263
wenzelm
parents: 21256
diff changeset
   448
  apply (subst power_abs [symmetric])
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   449
  apply (simp add: zero_le_even_power)
21263
wenzelm
parents: 21256
diff changeset
   450
  done
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   451
21263
wenzelm
parents: 21256
diff changeset
   452
lemma power_minus_even [simp]: "even n ==>
31017
2c227493ea56 stripped class recpower further
haftmann
parents: 30738
diff changeset
   453
    (- x)^n = (x^n::'a::{comm_ring_1})"
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   454
  apply (subst power_minus)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   455
  apply simp
21263
wenzelm
parents: 21256
diff changeset
   456
  done
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   457
21263
wenzelm
parents: 21256
diff changeset
   458
lemma power_minus_odd [simp]: "odd n ==>
31017
2c227493ea56 stripped class recpower further
haftmann
parents: 30738
diff changeset
   459
    (- x)^n = - (x^n::'a::{comm_ring_1})"
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   460
  apply (subst power_minus)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   461
  apply simp
21263
wenzelm
parents: 21256
diff changeset
   462
  done
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   463
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 33358
diff changeset
   464
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
   465
  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
   466
  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
   467
proof -
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29654
diff changeset
   468
  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
   469
  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
   470
  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
   471
  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
   472
qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29654
diff changeset
   473
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29654
diff changeset
   474
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
   475
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 33358
diff changeset
   476
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
   477
  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
   478
  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
   479
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
   480
  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
   481
  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
   482
  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
   483
next
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29654
diff changeset
   484
  case False
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29654
diff changeset
   485
  show ?thesis
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29654
diff changeset
   486
  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
   487
    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
   488
    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
   489
    moreover
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29654
diff changeset
   490
    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
   491
    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
   492
    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
   493
  next
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29654
diff changeset
   494
    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
   495
    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
   496
  qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29654
diff changeset
   497
qed
21263
wenzelm
parents: 21256
diff changeset
   498
58688
ddd124805260 restructured
haftmann
parents: 58687
diff changeset
   499
lemma (in linordered_idom) zero_le_power_iff [presburger]:
ddd124805260 restructured
haftmann
parents: 58687
diff changeset
   500
  "0 \<le> a ^ n \<longleftrightarrow> 0 \<le> a \<or> even n"
ddd124805260 restructured
haftmann
parents: 58687
diff changeset
   501
proof (cases "even n")
ddd124805260 restructured
haftmann
parents: 58687
diff changeset
   502
  case True
ddd124805260 restructured
haftmann
parents: 58687
diff changeset
   503
  then have "2 dvd n" by (simp add: even_def)
ddd124805260 restructured
haftmann
parents: 58687
diff changeset
   504
  then obtain k where "n = 2 * k" ..
ddd124805260 restructured
haftmann
parents: 58687
diff changeset
   505
  thus ?thesis by (simp add: zero_le_even_power True)
ddd124805260 restructured
haftmann
parents: 58687
diff changeset
   506
next
ddd124805260 restructured
haftmann
parents: 58687
diff changeset
   507
  case False
ddd124805260 restructured
haftmann
parents: 58687
diff changeset
   508
  then obtain k where "n = 2 * k + 1" ..
ddd124805260 restructured
haftmann
parents: 58687
diff changeset
   509
  moreover have "a ^ (2 * k) \<le> 0 \<Longrightarrow> a = 0"
ddd124805260 restructured
haftmann
parents: 58687
diff changeset
   510
    by (induct k) (auto simp add: zero_le_mult_iff mult_le_0_iff)
ddd124805260 restructured
haftmann
parents: 58687
diff changeset
   511
  ultimately show ?thesis
ddd124805260 restructured
haftmann
parents: 58687
diff changeset
   512
    by (auto simp add: zero_le_mult_iff zero_le_even_power)
ddd124805260 restructured
haftmann
parents: 58687
diff changeset
   513
qed
25600
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   514
21263
wenzelm
parents: 21256
diff changeset
   515
text {* Simplify, when the exponent is a numeral *}
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   516
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 45607
diff changeset
   517
lemmas zero_le_power_eq_numeral [simp] =
54227
63b441f49645 moving generic lemmas out of theory parity, disregarding some unused auxiliary lemmas;
haftmann
parents: 47225
diff changeset
   518
  zero_le_power_eq [of _ "numeral w"] for w
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   519
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 45607
diff changeset
   520
lemmas zero_less_power_eq_numeral [simp] =
54227
63b441f49645 moving generic lemmas out of theory parity, disregarding some unused auxiliary lemmas;
haftmann
parents: 47225
diff changeset
   521
  zero_less_power_eq [of _ "numeral w"] for w
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   522
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 45607
diff changeset
   523
lemmas power_le_zero_eq_numeral [simp] =
54227
63b441f49645 moving generic lemmas out of theory parity, disregarding some unused auxiliary lemmas;
haftmann
parents: 47225
diff changeset
   524
  power_le_zero_eq [of _ "numeral w"] for w
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   525
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 45607
diff changeset
   526
lemmas power_less_zero_eq_numeral [simp] =
54227
63b441f49645 moving generic lemmas out of theory parity, disregarding some unused auxiliary lemmas;
haftmann
parents: 47225
diff changeset
   527
  power_less_zero_eq [of _ "numeral w"] for w
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   528
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 45607
diff changeset
   529
lemmas zero_less_power_nat_eq_numeral [simp] =
54227
63b441f49645 moving generic lemmas out of theory parity, disregarding some unused auxiliary lemmas;
haftmann
parents: 47225
diff changeset
   530
  nat_zero_less_power_iff [of _ "numeral w"] for w
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   531
54227
63b441f49645 moving generic lemmas out of theory parity, disregarding some unused auxiliary lemmas;
haftmann
parents: 47225
diff changeset
   532
lemmas power_eq_0_iff_numeral [simp] =
63b441f49645 moving generic lemmas out of theory parity, disregarding some unused auxiliary lemmas;
haftmann
parents: 47225
diff changeset
   533
  power_eq_0_iff [of _ "numeral w"] for w
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   534
54227
63b441f49645 moving generic lemmas out of theory parity, disregarding some unused auxiliary lemmas;
haftmann
parents: 47225
diff changeset
   535
lemmas power_even_abs_numeral [simp] =
63b441f49645 moving generic lemmas out of theory parity, disregarding some unused auxiliary lemmas;
haftmann
parents: 47225
diff changeset
   536
  power_even_abs [of "numeral w" _] for w
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   537
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   538
end
54227
63b441f49645 moving generic lemmas out of theory parity, disregarding some unused auxiliary lemmas;
haftmann
parents: 47225
diff changeset
   539