src/HOL/Integ/Parity.thy
author haftmann
Mon, 06 Jun 2005 15:09:47 +0200
changeset 16303 fee0a02f61bb
parent 15251 bb6f072c8d10
child 16413 47ffc49c7d7b
permissions -rw-r--r--
integrated MacOS X installation instructions
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
     1
(*  Title:      Parity.thy
14450
3d2529f48b07 new thm
paulson
parents: 14443
diff changeset
     2
    ID:         $Id$
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
     3
    Author:     Jeremy Avigad
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
     4
*)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
     5
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
     6
header {* Parity: Even and Odd for ints and nats*}
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
     7
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15003
diff changeset
     8
theory Parity
15140
322485b816ac import -> imports
nipkow
parents: 15131
diff changeset
     9
imports Divides IntDiv NatSimprocs
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15003
diff changeset
    10
begin
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    11
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    12
axclass even_odd < type
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    13
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    14
instance int :: even_odd ..
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    15
instance nat :: even_odd ..
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    16
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    17
consts
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    18
  even :: "'a::even_odd => bool"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    19
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    20
syntax 
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    21
  odd :: "'a::even_odd => bool"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    22
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    23
translations 
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    24
  "odd x" == "~even x" 
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    25
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    26
defs (overloaded)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    27
  even_def: "even (x::int) == x mod 2 = 0"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    28
  even_nat_def: "even (x::nat) == even (int x)"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    29
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    30
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    31
subsection {* Casting a nat power to an integer *}
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    32
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    33
lemma zpow_int: "int (x^y) = (int x)^y"
15251
bb6f072c8d10 converted some induct_tac to induct
paulson
parents: 15140
diff changeset
    34
  apply (induct y)
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    35
  apply (simp, simp add: zmult_int [THEN sym])
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    36
  done
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    37
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    38
subsection {* Even and odd are mutually exclusive *}
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    39
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    40
lemma int_pos_lt_two_imp_zero_or_one: 
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    41
    "0 <= x ==> (x::int) < 2 ==> x = 0 | x = 1"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    42
  by auto
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    43
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    44
lemma neq_one_mod_two [simp]: "((x::int) mod 2 ~= 0) = (x mod 2 = 1)"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    45
  apply (subgoal_tac "x mod 2 = 0 | x mod 2 = 1", force)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    46
  apply (rule int_pos_lt_two_imp_zero_or_one, auto)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    47
  done
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    48
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    49
subsection {* Behavior under integer arithmetic operations *}
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    50
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    51
lemma even_times_anything: "even (x::int) ==> even (x * y)"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    52
  by (simp add: even_def zmod_zmult1_eq')
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    53
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    54
lemma anything_times_even: "even (y::int) ==> even (x * y)"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    55
  by (simp add: even_def zmod_zmult1_eq)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    56
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    57
lemma odd_times_odd: "odd (x::int) ==> odd y ==> odd (x * y)"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    58
  by (simp add: even_def zmod_zmult1_eq)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    59
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    60
lemma even_product: "even((x::int) * y) = (even x | even y)"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    61
  apply (auto simp add: even_times_anything anything_times_even) 
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    62
  apply (rule ccontr)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    63
  apply (auto simp add: odd_times_odd)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    64
  done
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    65
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    66
lemma even_plus_even: "even (x::int) ==> even y ==> even (x + y)"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    67
  by (simp add: even_def zmod_zadd1_eq)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    68
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    69
lemma even_plus_odd: "even (x::int) ==> odd y ==> odd (x + y)"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    70
  by (simp add: even_def zmod_zadd1_eq)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    71
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    72
lemma odd_plus_even: "odd (x::int) ==> even y ==> odd (x + y)"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    73
  by (simp add: even_def zmod_zadd1_eq)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    74
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    75
lemma odd_plus_odd: "odd (x::int) ==> odd y ==> even (x + y)"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    76
  by (simp add: even_def zmod_zadd1_eq)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    77
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    78
lemma even_sum: "even ((x::int) + y) = ((even x & even y) | (odd x & odd y))"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    79
  apply (auto intro: even_plus_even odd_plus_odd)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    80
  apply (rule ccontr, simp add: even_plus_odd)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    81
  apply (rule ccontr, simp add: odd_plus_even)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    82
  done
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    83
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    84
lemma even_neg: "even (-(x::int)) = even x"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    85
  by (auto simp add: even_def zmod_zminus1_eq_if)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    86
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    87
lemma even_difference: 
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    88
  "even ((x::int) - y) = ((even x & even y) | (odd x & odd y))"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    89
  by (simp only: diff_minus even_sum even_neg)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    90
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    91
lemma even_pow_gt_zero [rule_format]: 
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    92
    "even (x::int) ==> 0 < n --> even (x^n)"
15251
bb6f072c8d10 converted some induct_tac to induct
paulson
parents: 15140
diff changeset
    93
  apply (induct n)
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    94
  apply (auto simp add: even_product)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    95
  done
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    96
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    97
lemma odd_pow: "odd x ==> odd((x::int)^n)"
15251
bb6f072c8d10 converted some induct_tac to induct
paulson
parents: 15140
diff changeset
    98
  apply (induct n)
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
    99
  apply (simp add: even_def)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   100
  apply (simp add: even_product)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   101
  done
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   102
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   103
lemma even_power: "even ((x::int)^n) = (even x & 0 < n)"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   104
  apply (auto simp add: even_pow_gt_zero) 
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   105
  apply (erule contrapos_pp, erule odd_pow)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   106
  apply (erule contrapos_pp, simp add: even_def)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   107
  done
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   108
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   109
lemma even_zero: "even (0::int)"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   110
  by (simp add: even_def)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   111
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   112
lemma odd_one: "odd (1::int)"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   113
  by (simp add: even_def)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   114
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   115
lemmas even_odd_simps [simp] = even_def[of "number_of v",standard] even_zero 
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   116
  odd_one even_product even_sum even_neg even_difference even_power
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   117
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   118
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   119
subsection {* Equivalent definitions *}
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   120
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   121
lemma two_times_even_div_two: "even (x::int) ==> 2 * (x div 2) = x" 
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   122
  by (auto simp add: even_def)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   123
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   124
lemma two_times_odd_div_two_plus_one: "odd (x::int) ==> 
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   125
    2 * (x div 2) + 1 = x"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   126
  apply (insert zmod_zdiv_equality [of x 2, THEN sym])
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   127
  by (simp add: even_def)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   128
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   129
lemma even_equiv_def: "even (x::int) = (EX y. x = 2 * y)"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   130
  apply auto
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   131
  apply (rule exI)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   132
  by (erule two_times_even_div_two [THEN sym])
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   133
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   134
lemma odd_equiv_def: "odd (x::int) = (EX y. x = 2 * y + 1)"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   135
  apply auto
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   136
  apply (rule exI)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   137
  by (erule two_times_odd_div_two_plus_one [THEN sym])
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   138
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   139
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   140
subsection {* even and odd for nats *}
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   141
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   142
lemma pos_int_even_equiv_nat_even: "0 \<le> x ==> even x = even (nat x)"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   143
  by (simp add: even_nat_def)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   144
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   145
lemma even_nat_product: "even((x::nat) * y) = (even x | even y)"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   146
  by (simp add: even_nat_def zmult_int [THEN sym])
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   147
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   148
lemma even_nat_sum: "even ((x::nat) + y) = 
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   149
    ((even x & even y) | (odd x & odd y))"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   150
  by (unfold even_nat_def, simp)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   151
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   152
lemma even_nat_difference: 
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   153
    "even ((x::nat) - y) = (x < y | (even x & even y) | (odd x & odd y))"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   154
  apply (auto simp add: even_nat_def zdiff_int [THEN sym])
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   155
  apply (case_tac "x < y", auto simp add: zdiff_int [THEN sym])
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   156
  apply (case_tac "x < y", auto simp add: zdiff_int [THEN sym])
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   157
  done
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   158
14436
77017c49c004 some new results
paulson
parents: 14430
diff changeset
   159
lemma even_nat_Suc: "even (Suc x) = odd x"
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   160
  by (simp add: even_nat_def)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   161
14436
77017c49c004 some new results
paulson
parents: 14430
diff changeset
   162
text{*Compatibility, in case Avigad uses this*}
77017c49c004 some new results
paulson
parents: 14430
diff changeset
   163
lemmas even_nat_suc = even_nat_Suc
77017c49c004 some new results
paulson
parents: 14430
diff changeset
   164
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   165
lemma even_nat_power: "even ((x::nat)^y) = (even x & 0 < y)"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   166
  by (simp add: even_nat_def zpow_int)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   167
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   168
lemma even_nat_zero: "even (0::nat)"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   169
  by (simp add: even_nat_def)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   170
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   171
lemmas even_odd_nat_simps [simp] = even_nat_def[of "number_of v",standard] 
14436
77017c49c004 some new results
paulson
parents: 14430
diff changeset
   172
  even_nat_zero even_nat_Suc even_nat_product even_nat_sum even_nat_power
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   173
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   174
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   175
subsection {* Equivalent definitions *}
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   176
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   177
lemma nat_lt_two_imp_zero_or_one: "(x::nat) < Suc (Suc 0) ==> 
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   178
    x = 0 | x = Suc 0"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   179
  by auto
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   180
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   181
lemma even_nat_mod_two_eq_zero: "even (x::nat) ==> x mod (Suc (Suc 0)) = 0"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   182
  apply (insert mod_div_equality [of x "Suc (Suc 0)", THEN sym])
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   183
  apply (drule subst, assumption)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   184
  apply (subgoal_tac "x mod Suc (Suc 0) = 0 | x mod Suc (Suc 0) = Suc 0")
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   185
  apply force
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   186
  apply (subgoal_tac "0 < Suc (Suc 0)")
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   187
  apply (frule mod_less_divisor [of "Suc (Suc 0)" x])
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   188
  apply (erule nat_lt_two_imp_zero_or_one, auto)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   189
  done
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   190
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   191
lemma odd_nat_mod_two_eq_one: "odd (x::nat) ==> x mod (Suc (Suc 0)) = Suc 0"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   192
  apply (insert mod_div_equality [of x "Suc (Suc 0)", THEN sym])
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   193
  apply (drule subst, assumption)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   194
  apply (subgoal_tac "x mod Suc (Suc 0) = 0 | x mod Suc (Suc 0) = Suc 0")
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   195
  apply force 
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   196
  apply (subgoal_tac "0 < Suc (Suc 0)")
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   197
  apply (frule mod_less_divisor [of "Suc (Suc 0)" x])
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   198
  apply (erule nat_lt_two_imp_zero_or_one, auto)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   199
  done
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   200
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   201
lemma even_nat_equiv_def: "even (x::nat) = (x mod Suc (Suc 0) = 0)" 
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   202
  apply (rule iffI)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   203
  apply (erule even_nat_mod_two_eq_zero)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   204
  apply (insert odd_nat_mod_two_eq_one [of x], auto)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   205
  done
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   206
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   207
lemma odd_nat_equiv_def: "odd (x::nat) = (x mod Suc (Suc 0) = Suc 0)"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   208
  apply (auto simp add: even_nat_equiv_def)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   209
  apply (subgoal_tac "x mod (Suc (Suc 0)) < Suc (Suc 0)")
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   210
  apply (frule nat_lt_two_imp_zero_or_one, auto)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   211
  done
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   212
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   213
lemma even_nat_div_two_times_two: "even (x::nat) ==> 
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   214
    Suc (Suc 0) * (x div Suc (Suc 0)) = x"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   215
  apply (insert mod_div_equality [of x "Suc (Suc 0)", THEN sym])
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   216
  apply (drule even_nat_mod_two_eq_zero, simp)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   217
  done
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   218
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   219
lemma odd_nat_div_two_times_two_plus_one: "odd (x::nat) ==> 
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   220
    Suc( Suc (Suc 0) * (x div Suc (Suc 0))) = x"  
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   221
  apply (insert mod_div_equality [of x "Suc (Suc 0)", THEN sym])
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   222
  apply (drule odd_nat_mod_two_eq_one, simp)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   223
  done
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   224
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   225
lemma even_nat_equiv_def2: "even (x::nat) = (EX y. x = Suc (Suc 0) * y)"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   226
  apply (rule iffI, rule exI)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   227
  apply (erule even_nat_div_two_times_two [THEN sym], auto)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   228
  done
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   229
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   230
lemma odd_nat_equiv_def2: "odd (x::nat) = (EX y. x = Suc(Suc (Suc 0) * y))"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   231
  apply (rule iffI, rule exI)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   232
  apply (erule odd_nat_div_two_times_two_plus_one [THEN sym], auto)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   233
  done
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   234
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   235
subsection {* Powers of negative one *}
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   236
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   237
lemma neg_one_even_odd_power:
15003
6145dd7538d7 replaced monomorphic abs definitions by abs_if
paulson
parents: 14981
diff changeset
   238
     "(even x --> (-1::'a::{number_ring,recpower})^x = 1) & 
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   239
      (odd x --> (-1::'a)^x = -1)"
15251
bb6f072c8d10 converted some induct_tac to induct
paulson
parents: 15140
diff changeset
   240
  apply (induct x)
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   241
  apply (simp, simp add: power_Suc)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   242
  done
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   243
14436
77017c49c004 some new results
paulson
parents: 14430
diff changeset
   244
lemma neg_one_even_power [simp]:
15003
6145dd7538d7 replaced monomorphic abs definitions by abs_if
paulson
parents: 14981
diff changeset
   245
     "even x ==> (-1::'a::{number_ring,recpower})^x = 1"
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   246
  by (rule neg_one_even_odd_power [THEN conjunct1, THEN mp], assumption)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   247
14436
77017c49c004 some new results
paulson
parents: 14430
diff changeset
   248
lemma neg_one_odd_power [simp]:
15003
6145dd7538d7 replaced monomorphic abs definitions by abs_if
paulson
parents: 14981
diff changeset
   249
     "odd x ==> (-1::'a::{number_ring,recpower})^x = -1"
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   250
  by (rule neg_one_even_odd_power [THEN conjunct2, THEN mp], assumption)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   251
14443
75910c7557c5 generic theorems about exponentials; general tidying up
paulson
parents: 14436
diff changeset
   252
lemma neg_power_if:
15003
6145dd7538d7 replaced monomorphic abs definitions by abs_if
paulson
parents: 14981
diff changeset
   253
     "(-x::'a::{comm_ring_1,recpower}) ^ n = 
14443
75910c7557c5 generic theorems about exponentials; general tidying up
paulson
parents: 14436
diff changeset
   254
      (if even n then (x ^ n) else -(x ^ n))"
75910c7557c5 generic theorems about exponentials; general tidying up
paulson
parents: 14436
diff changeset
   255
  by (induct n, simp_all split: split_if_asm add: power_Suc) 
75910c7557c5 generic theorems about exponentials; general tidying up
paulson
parents: 14436
diff changeset
   256
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   257
14450
3d2529f48b07 new thm
paulson
parents: 14443
diff changeset
   258
subsection {* An Equivalence for @{term "0 \<le> a^n"} *}
3d2529f48b07 new thm
paulson
parents: 14443
diff changeset
   259
3d2529f48b07 new thm
paulson
parents: 14443
diff changeset
   260
lemma even_power_le_0_imp_0:
15003
6145dd7538d7 replaced monomorphic abs definitions by abs_if
paulson
parents: 14981
diff changeset
   261
     "a ^ (2*k) \<le> (0::'a::{ordered_idom,recpower}) ==> a=0"
14450
3d2529f48b07 new thm
paulson
parents: 14443
diff changeset
   262
apply (induct k) 
3d2529f48b07 new thm
paulson
parents: 14443
diff changeset
   263
apply (auto simp add: zero_le_mult_iff mult_le_0_iff power_Suc)  
3d2529f48b07 new thm
paulson
parents: 14443
diff changeset
   264
done
3d2529f48b07 new thm
paulson
parents: 14443
diff changeset
   265
3d2529f48b07 new thm
paulson
parents: 14443
diff changeset
   266
lemma zero_le_power_iff:
15003
6145dd7538d7 replaced monomorphic abs definitions by abs_if
paulson
parents: 14981
diff changeset
   267
     "(0 \<le> a^n) = (0 \<le> (a::'a::{ordered_idom,recpower}) | even n)"
14450
3d2529f48b07 new thm
paulson
parents: 14443
diff changeset
   268
      (is "?P n")
3d2529f48b07 new thm
paulson
parents: 14443
diff changeset
   269
proof cases
3d2529f48b07 new thm
paulson
parents: 14443
diff changeset
   270
  assume even: "even n"
14473
846c237bd9b3 stylistic tweaks
paulson
parents: 14450
diff changeset
   271
  then obtain k where "n = 2*k"
14450
3d2529f48b07 new thm
paulson
parents: 14443
diff changeset
   272
    by (auto simp add: even_nat_equiv_def2 numeral_2_eq_2)
3d2529f48b07 new thm
paulson
parents: 14443
diff changeset
   273
  thus ?thesis by (simp add: zero_le_even_power even) 
3d2529f48b07 new thm
paulson
parents: 14443
diff changeset
   274
next
3d2529f48b07 new thm
paulson
parents: 14443
diff changeset
   275
  assume odd: "odd n"
14473
846c237bd9b3 stylistic tweaks
paulson
parents: 14450
diff changeset
   276
  then obtain k where "n = Suc(2*k)"
14450
3d2529f48b07 new thm
paulson
parents: 14443
diff changeset
   277
    by (auto simp add: odd_nat_equiv_def2 numeral_2_eq_2)
3d2529f48b07 new thm
paulson
parents: 14443
diff changeset
   278
  thus ?thesis
3d2529f48b07 new thm
paulson
parents: 14443
diff changeset
   279
    by (auto simp add: power_Suc zero_le_mult_iff zero_le_even_power 
3d2529f48b07 new thm
paulson
parents: 14443
diff changeset
   280
             dest!: even_power_le_0_imp_0) 
3d2529f48b07 new thm
paulson
parents: 14443
diff changeset
   281
qed 
3d2529f48b07 new thm
paulson
parents: 14443
diff changeset
   282
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   283
subsection {* Miscellaneous *}
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   284
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   285
lemma even_plus_one_div_two: "even (x::int) ==> (x + 1) div 2 = x div 2"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   286
  apply (subst zdiv_zadd1_eq)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   287
  apply (simp add: even_def)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   288
  done
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   289
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   290
lemma odd_plus_one_div_two: "odd (x::int) ==> (x + 1) div 2 = x div 2 + 1"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   291
  apply (subst zdiv_zadd1_eq)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   292
  apply (simp add: even_def)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   293
  done
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   294
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   295
lemma div_Suc: "Suc a div c = a div c + Suc 0 div c + 
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   296
    (a mod c + Suc 0 mod c) div c"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   297
  apply (subgoal_tac "Suc a = a + Suc 0")
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   298
  apply (erule ssubst)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   299
  apply (rule div_add1_eq, simp)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   300
  done
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   301
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   302
lemma even_nat_plus_one_div_two: "even (x::nat) ==> 
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   303
   (Suc x) div Suc (Suc 0) = x div Suc (Suc 0)"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   304
  apply (subst div_Suc)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   305
  apply (simp add: even_nat_equiv_def)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   306
  done
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   307
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   308
lemma odd_nat_plus_one_div_two: "odd (x::nat) ==> 
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   309
    (Suc x) div Suc (Suc 0) = Suc (x div Suc (Suc 0))"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   310
  apply (subst div_Suc)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   311
  apply (simp add: odd_nat_equiv_def)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   312
  done
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   313
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents:
diff changeset
   314
end