wenzelm@21263: (* Title: HOL/Library/Parity.thy wenzelm@21256: ID: $Id$ haftmann@25600: Author: Jeremy Avigad, Jacques D. Fleuriot wenzelm@21256: *) wenzelm@21256: wenzelm@21256: header {* Even and Odd for int and nat *} wenzelm@21256: wenzelm@21256: theory Parity haftmann@25594: imports PreList wenzelm@21256: begin wenzelm@21256: haftmann@22473: class even_odd = type + haftmann@22390: fixes even :: "'a \ bool" wenzelm@21256: wenzelm@21256: abbreviation haftmann@22390: odd :: "'a\even_odd \ bool" where haftmann@22390: "odd x \ \ even x" haftmann@22390: haftmann@25571: instantiation int and nat :: even_odd haftmann@25571: begin haftmann@25571: haftmann@25571: definition haftmann@25571: even_def [presburger]: "even x \ (x\int) mod 2 = 0" haftmann@22390: haftmann@25571: definition haftmann@25571: even_nat_def [presburger]: "even x \ even (int x)" haftmann@25571: haftmann@25571: instance .. haftmann@25571: haftmann@25571: end wenzelm@21256: wenzelm@21256: wenzelm@21256: subsection {* Even and odd are mutually exclusive *} wenzelm@21256: wenzelm@21263: lemma int_pos_lt_two_imp_zero_or_one: wenzelm@21256: "0 <= x ==> (x::int) < 2 ==> x = 0 | x = 1" chaieb@23522: by presburger wenzelm@21256: chaieb@23522: lemma neq_one_mod_two [simp, presburger]: chaieb@23522: "((x::int) mod 2 ~= 0) = (x mod 2 = 1)" by presburger wenzelm@21256: haftmann@25600: wenzelm@21256: subsection {* Behavior under integer arithmetic operations *} wenzelm@21256: wenzelm@21256: lemma even_times_anything: "even (x::int) ==> even (x * y)" wenzelm@21256: by (simp add: even_def zmod_zmult1_eq') wenzelm@21256: wenzelm@21256: lemma anything_times_even: "even (y::int) ==> even (x * y)" wenzelm@21256: by (simp add: even_def zmod_zmult1_eq) wenzelm@21256: wenzelm@21256: lemma odd_times_odd: "odd (x::int) ==> odd y ==> odd (x * y)" wenzelm@21256: by (simp add: even_def zmod_zmult1_eq) wenzelm@21256: chaieb@23522: lemma even_product[presburger]: "even((x::int) * y) = (even x | even y)" wenzelm@21263: apply (auto simp add: even_times_anything anything_times_even) wenzelm@21256: apply (rule ccontr) wenzelm@21256: apply (auto simp add: odd_times_odd) wenzelm@21256: done wenzelm@21256: wenzelm@21256: lemma even_plus_even: "even (x::int) ==> even y ==> even (x + y)" chaieb@23522: by presburger wenzelm@21256: wenzelm@21256: lemma even_plus_odd: "even (x::int) ==> odd y ==> odd (x + y)" chaieb@23522: by presburger wenzelm@21256: wenzelm@21256: lemma odd_plus_even: "odd (x::int) ==> even y ==> odd (x + y)" chaieb@23522: by presburger wenzelm@21256: chaieb@23522: lemma odd_plus_odd: "odd (x::int) ==> odd y ==> even (x + y)" by presburger wenzelm@21256: chaieb@23522: lemma even_sum[presburger]: "even ((x::int) + y) = ((even x & even y) | (odd x & odd y))" chaieb@23522: by presburger wenzelm@21256: chaieb@23522: lemma even_neg[presburger]: "even (-(x::int)) = even x" by presburger wenzelm@21256: wenzelm@21263: lemma even_difference: chaieb@23522: "even ((x::int) - y) = ((even x & even y) | (odd x & odd y))" by presburger wenzelm@21256: wenzelm@21263: lemma even_pow_gt_zero: wenzelm@21263: "even (x::int) ==> 0 < n ==> even (x^n)" wenzelm@21263: by (induct n) (auto simp add: even_product) wenzelm@21256: chaieb@23522: lemma odd_pow_iff[presburger]: "odd ((x::int) ^ n) \ (n = 0 \ odd x)" chaieb@23522: apply (induct n, simp_all) chaieb@23522: apply presburger chaieb@23522: apply (case_tac n, auto) chaieb@23522: apply (simp_all add: even_product) wenzelm@21256: done wenzelm@21256: chaieb@23522: lemma odd_pow: "odd x ==> odd((x::int)^n)" by (simp add: odd_pow_iff) chaieb@23522: chaieb@23522: lemma even_power[presburger]: "even ((x::int)^n) = (even x & 0 < n)" wenzelm@21263: apply (auto simp add: even_pow_gt_zero) wenzelm@21256: apply (erule contrapos_pp, erule odd_pow) wenzelm@21256: apply (erule contrapos_pp, simp add: even_def) wenzelm@21256: done wenzelm@21256: chaieb@23522: lemma even_zero[presburger]: "even (0::int)" by presburger wenzelm@21256: chaieb@23522: lemma odd_one[presburger]: "odd (1::int)" by presburger wenzelm@21256: wenzelm@21263: lemmas even_odd_simps [simp] = even_def[of "number_of v",standard] even_zero wenzelm@21256: odd_one even_product even_sum even_neg even_difference even_power wenzelm@21256: wenzelm@21256: wenzelm@21256: subsection {* Equivalent definitions *} wenzelm@21256: chaieb@23522: lemma two_times_even_div_two: "even (x::int) ==> 2 * (x div 2) = x" chaieb@23522: by presburger wenzelm@21256: wenzelm@21263: lemma two_times_odd_div_two_plus_one: "odd (x::int) ==> chaieb@23522: 2 * (x div 2) + 1 = x" by presburger wenzelm@21256: chaieb@23522: lemma even_equiv_def: "even (x::int) = (EX y. x = 2 * y)" by presburger wenzelm@21256: chaieb@23522: lemma odd_equiv_def: "odd (x::int) = (EX y. x = 2 * y + 1)" by presburger wenzelm@21256: wenzelm@21256: subsection {* even and odd for nats *} wenzelm@21256: wenzelm@21256: lemma pos_int_even_equiv_nat_even: "0 \ x ==> even x = even (nat x)" wenzelm@21256: by (simp add: even_nat_def) wenzelm@21256: chaieb@23522: lemma even_nat_product[presburger]: "even((x::nat) * y) = (even x | even y)" huffman@23431: by (simp add: even_nat_def int_mult) wenzelm@21256: chaieb@23522: lemma even_nat_sum[presburger]: "even ((x::nat) + y) = chaieb@23522: ((even x & even y) | (odd x & odd y))" by presburger wenzelm@21256: chaieb@23522: lemma even_nat_difference[presburger]: wenzelm@21256: "even ((x::nat) - y) = (x < y | (even x & even y) | (odd x & odd y))" chaieb@23522: by presburger wenzelm@21256: chaieb@23522: lemma even_nat_Suc[presburger]: "even (Suc x) = odd x" by presburger wenzelm@21256: chaieb@23522: lemma even_nat_power[presburger]: "even ((x::nat)^y) = (even x & 0 < y)" huffman@23431: by (simp add: even_nat_def int_power) wenzelm@21256: chaieb@23522: lemma even_nat_zero[presburger]: "even (0::nat)" by presburger wenzelm@21256: wenzelm@21263: lemmas even_odd_nat_simps [simp] = even_nat_def[of "number_of v",standard] wenzelm@21256: even_nat_zero even_nat_Suc even_nat_product even_nat_sum even_nat_power wenzelm@21256: wenzelm@21256: wenzelm@21256: subsection {* Equivalent definitions *} wenzelm@21256: wenzelm@21263: lemma nat_lt_two_imp_zero_or_one: "(x::nat) < Suc (Suc 0) ==> chaieb@23522: x = 0 | x = Suc 0" by presburger wenzelm@21256: wenzelm@21256: lemma even_nat_mod_two_eq_zero: "even (x::nat) ==> x mod (Suc (Suc 0)) = 0" chaieb@23522: by presburger wenzelm@21256: wenzelm@21256: lemma odd_nat_mod_two_eq_one: "odd (x::nat) ==> x mod (Suc (Suc 0)) = Suc 0" chaieb@23522: by presburger wenzelm@21256: wenzelm@21263: lemma even_nat_equiv_def: "even (x::nat) = (x mod Suc (Suc 0) = 0)" chaieb@23522: by presburger wenzelm@21256: wenzelm@21256: lemma odd_nat_equiv_def: "odd (x::nat) = (x mod Suc (Suc 0) = Suc 0)" chaieb@23522: by presburger wenzelm@21256: wenzelm@21263: lemma even_nat_div_two_times_two: "even (x::nat) ==> chaieb@23522: Suc (Suc 0) * (x div Suc (Suc 0)) = x" by presburger wenzelm@21256: wenzelm@21263: lemma odd_nat_div_two_times_two_plus_one: "odd (x::nat) ==> chaieb@23522: Suc( Suc (Suc 0) * (x div Suc (Suc 0))) = x" by presburger wenzelm@21256: wenzelm@21256: lemma even_nat_equiv_def2: "even (x::nat) = (EX y. x = Suc (Suc 0) * y)" chaieb@23522: by presburger wenzelm@21256: wenzelm@21256: lemma odd_nat_equiv_def2: "odd (x::nat) = (EX y. x = Suc(Suc (Suc 0) * y))" chaieb@23522: by presburger wenzelm@21256: haftmann@25600: wenzelm@21256: subsection {* Parity and powers *} wenzelm@21256: wenzelm@21263: lemma minus_one_even_odd_power: wenzelm@21263: "(even x --> (- 1::'a::{comm_ring_1,recpower})^x = 1) & wenzelm@21256: (odd x --> (- 1::'a)^x = - 1)" wenzelm@21256: apply (induct x) wenzelm@21256: apply (rule conjI) wenzelm@21256: apply simp wenzelm@21256: apply (insert even_nat_zero, blast) wenzelm@21256: apply (simp add: power_Suc) wenzelm@21263: done wenzelm@21256: wenzelm@21256: lemma minus_one_even_power [simp]: wenzelm@21263: "even x ==> (- 1::'a::{comm_ring_1,recpower})^x = 1" wenzelm@21263: using minus_one_even_odd_power by blast wenzelm@21256: wenzelm@21256: lemma minus_one_odd_power [simp]: wenzelm@21263: "odd x ==> (- 1::'a::{comm_ring_1,recpower})^x = - 1" wenzelm@21263: using minus_one_even_odd_power by blast wenzelm@21256: wenzelm@21256: lemma neg_one_even_odd_power: wenzelm@21263: "(even x --> (-1::'a::{number_ring,recpower})^x = 1) & wenzelm@21256: (odd x --> (-1::'a)^x = -1)" wenzelm@21256: apply (induct x) wenzelm@21256: apply (simp, simp add: power_Suc) wenzelm@21256: done wenzelm@21256: wenzelm@21256: lemma neg_one_even_power [simp]: wenzelm@21263: "even x ==> (-1::'a::{number_ring,recpower})^x = 1" wenzelm@21263: using neg_one_even_odd_power by blast wenzelm@21256: wenzelm@21256: lemma neg_one_odd_power [simp]: wenzelm@21263: "odd x ==> (-1::'a::{number_ring,recpower})^x = -1" wenzelm@21263: using neg_one_even_odd_power by blast wenzelm@21256: wenzelm@21256: lemma neg_power_if: wenzelm@21263: "(-x::'a::{comm_ring_1,recpower}) ^ n = wenzelm@21256: (if even n then (x ^ n) else -(x ^ n))" wenzelm@21263: apply (induct n) wenzelm@21263: apply (simp_all split: split_if_asm add: power_Suc) wenzelm@21263: done wenzelm@21256: wenzelm@21263: lemma zero_le_even_power: "even n ==> wenzelm@21256: 0 <= (x::'a::{recpower,ordered_ring_strict}) ^ n" wenzelm@21256: apply (simp add: even_nat_equiv_def2) wenzelm@21256: apply (erule exE) wenzelm@21256: apply (erule ssubst) wenzelm@21256: apply (subst power_add) wenzelm@21256: apply (rule zero_le_square) wenzelm@21256: done wenzelm@21256: wenzelm@21263: lemma zero_le_odd_power: "odd n ==> wenzelm@21256: (0 <= (x::'a::{recpower,ordered_idom}) ^ n) = (0 <= x)" wenzelm@21256: apply (simp add: odd_nat_equiv_def2) wenzelm@21256: apply (erule exE) wenzelm@21256: apply (erule ssubst) wenzelm@21256: apply (subst power_Suc) wenzelm@21256: apply (subst power_add) wenzelm@21256: apply (subst zero_le_mult_iff) wenzelm@21256: apply auto nipkow@25162: apply (subgoal_tac "x = 0 & y > 0") wenzelm@21256: apply (erule conjE, assumption) wenzelm@21263: apply (subst power_eq_0_iff [symmetric]) wenzelm@21256: apply (subgoal_tac "0 <= x^y * x^y") wenzelm@21256: apply simp wenzelm@21256: apply (rule zero_le_square)+ wenzelm@21263: done wenzelm@21256: chaieb@23522: lemma zero_le_power_eq[presburger]: "(0 <= (x::'a::{recpower,ordered_idom}) ^ n) = wenzelm@21256: (even n | (odd n & 0 <= x))" wenzelm@21256: apply auto wenzelm@21263: apply (subst zero_le_odd_power [symmetric]) wenzelm@21256: apply assumption+ wenzelm@21256: apply (erule zero_le_even_power) wenzelm@21263: apply (subst zero_le_odd_power) wenzelm@21256: apply assumption+ wenzelm@21263: done wenzelm@21256: chaieb@23522: lemma zero_less_power_eq[presburger]: "(0 < (x::'a::{recpower,ordered_idom}) ^ n) = wenzelm@21256: (n = 0 | (even n & x ~= 0) | (odd n & 0 < x))" wenzelm@21256: apply (rule iffI) wenzelm@21256: apply clarsimp wenzelm@21256: apply (rule conjI) wenzelm@21256: apply clarsimp wenzelm@21256: apply (rule ccontr) wenzelm@21256: apply (subgoal_tac "~ (0 <= x^n)") wenzelm@21256: apply simp wenzelm@21256: apply (subst zero_le_odd_power) wenzelm@21263: apply assumption wenzelm@21256: apply simp wenzelm@21256: apply (rule notI) wenzelm@21256: apply (simp add: power_0_left) wenzelm@21256: apply (rule notI) wenzelm@21256: apply (simp add: power_0_left) wenzelm@21256: apply auto wenzelm@21256: apply (subgoal_tac "0 <= x^n") wenzelm@21256: apply (frule order_le_imp_less_or_eq) wenzelm@21256: apply simp wenzelm@21256: apply (erule zero_le_even_power) wenzelm@21256: apply (subgoal_tac "0 <= x^n") wenzelm@21256: apply (frule order_le_imp_less_or_eq) wenzelm@21256: apply auto wenzelm@21256: apply (subst zero_le_odd_power) wenzelm@21256: apply assumption wenzelm@21256: apply (erule order_less_imp_le) wenzelm@21263: done wenzelm@21256: chaieb@23522: lemma power_less_zero_eq[presburger]: "((x::'a::{recpower,ordered_idom}) ^ n < 0) = chaieb@23522: (odd n & x < 0)" wenzelm@21263: apply (subst linorder_not_le [symmetric])+ wenzelm@21256: apply (subst zero_le_power_eq) wenzelm@21256: apply auto wenzelm@21263: done wenzelm@21256: chaieb@23522: lemma power_le_zero_eq[presburger]: "((x::'a::{recpower,ordered_idom}) ^ n <= 0) = wenzelm@21256: (n ~= 0 & ((odd n & x <= 0) | (even n & x = 0)))" wenzelm@21263: apply (subst linorder_not_less [symmetric])+ wenzelm@21256: apply (subst zero_less_power_eq) wenzelm@21256: apply auto wenzelm@21263: done wenzelm@21256: wenzelm@21263: lemma power_even_abs: "even n ==> wenzelm@21256: (abs (x::'a::{recpower,ordered_idom}))^n = x^n" wenzelm@21263: apply (subst power_abs [symmetric]) wenzelm@21256: apply (simp add: zero_le_even_power) wenzelm@21263: done wenzelm@21256: chaieb@23522: lemma zero_less_power_nat_eq[presburger]: "(0 < (x::nat) ^ n) = (n = 0 | 0 < x)" wenzelm@21263: by (induct n) auto wenzelm@21256: wenzelm@21263: lemma power_minus_even [simp]: "even n ==> wenzelm@21256: (- x)^n = (x^n::'a::{recpower,comm_ring_1})" wenzelm@21256: apply (subst power_minus) wenzelm@21256: apply simp wenzelm@21263: done wenzelm@21256: wenzelm@21263: lemma power_minus_odd [simp]: "odd n ==> wenzelm@21256: (- x)^n = - (x^n::'a::{recpower,comm_ring_1})" wenzelm@21256: apply (subst power_minus) wenzelm@21256: apply simp wenzelm@21263: done wenzelm@21256: wenzelm@21263: haftmann@25600: subsection {* General Lemmas About Division *} haftmann@25600: haftmann@25600: lemma Suc_times_mod_eq: "1 Suc (k * m) mod k = 1" haftmann@25600: apply (induct "m") haftmann@25600: apply (simp_all add: mod_Suc) haftmann@25600: done haftmann@25600: haftmann@25600: declare Suc_times_mod_eq [of "number_of w", standard, simp] haftmann@25600: haftmann@25600: lemma [simp]: "n div k \ (Suc n) div k" haftmann@25600: by (simp add: div_le_mono) haftmann@25600: haftmann@25600: lemma Suc_n_div_2_gt_zero [simp]: "(0::nat) < n ==> 0 < (n + 1) div 2" haftmann@25600: by arith haftmann@25600: haftmann@25600: lemma div_2_gt_zero [simp]: "(1::nat) < n ==> 0 < n div 2" haftmann@25600: by arith haftmann@25600: haftmann@25600: lemma mod_mult_self3 [simp]: "(k*n + m) mod n = m mod (n::nat)" haftmann@25600: by (simp add: mult_ac add_ac) haftmann@25600: haftmann@25600: lemma mod_mult_self4 [simp]: "Suc (k*n + m) mod n = Suc m mod n" haftmann@25600: proof - haftmann@25600: have "Suc (k * n + m) mod n = (k * n + Suc m) mod n" by simp haftmann@25600: also have "... = Suc m mod n" by (rule mod_mult_self3) haftmann@25600: finally show ?thesis . haftmann@25600: qed haftmann@25600: haftmann@25600: lemma mod_Suc_eq_Suc_mod: "Suc m mod n = Suc (m mod n) mod n" haftmann@25600: apply (subst mod_Suc [of m]) haftmann@25600: apply (subst mod_Suc [of "m mod n"], simp) haftmann@25600: done haftmann@25600: haftmann@25600: haftmann@25600: subsection {* More Even/Odd Results *} haftmann@25600: haftmann@25600: lemma even_mult_two_ex: "even(n) = (\m::nat. n = 2*m)" haftmann@25600: by (simp add: even_nat_equiv_def2 numeral_2_eq_2) haftmann@25600: haftmann@25600: lemma odd_Suc_mult_two_ex: "odd(n) = (\m. n = Suc (2*m))" haftmann@25600: by (simp add: odd_nat_equiv_def2 numeral_2_eq_2) haftmann@25600: haftmann@25600: lemma even_add [simp]: "even(m + n::nat) = (even m = even n)" haftmann@25600: by auto haftmann@25600: haftmann@25600: lemma odd_add [simp]: "odd(m + n::nat) = (odd m \ odd n)" haftmann@25600: by auto haftmann@25600: haftmann@25600: lemma div_Suc: "Suc a div c = a div c + Suc 0 div c + haftmann@25600: (a mod c + Suc 0 mod c) div c" haftmann@25600: apply (subgoal_tac "Suc a = a + Suc 0") haftmann@25600: apply (erule ssubst) haftmann@25600: apply (rule div_add1_eq, simp) haftmann@25600: done haftmann@25600: haftmann@25600: lemma lemma_even_div2 [simp]: "even (n::nat) ==> (n + 1) div 2 = n div 2" haftmann@25600: apply (simp add: numeral_2_eq_2) haftmann@25600: apply (subst div_Suc) haftmann@25600: apply (simp add: even_nat_mod_two_eq_zero) haftmann@25600: done haftmann@25600: haftmann@25600: lemma lemma_not_even_div2 [simp]: "~even n ==> (n + 1) div 2 = Suc (n div 2)" haftmann@25600: apply (simp add: numeral_2_eq_2) haftmann@25600: apply (subst div_Suc) haftmann@25600: apply (simp add: odd_nat_mod_two_eq_one) haftmann@25600: done haftmann@25600: haftmann@25600: lemma even_num_iff: "0 < n ==> even n = (~ even(n - 1 :: nat))" haftmann@25600: by (case_tac "n", auto) haftmann@25600: haftmann@25600: lemma even_even_mod_4_iff: "even (n::nat) = even (n mod 4)" haftmann@25600: apply (induct n, simp) haftmann@25600: apply (subst mod_Suc, simp) haftmann@25600: done haftmann@25600: haftmann@25600: lemma lemma_odd_mod_4_div_2: "n mod 4 = (3::nat) ==> odd((n - 1) div 2)" haftmann@25600: apply (rule_tac t = n and n1 = 4 in mod_div_equality [THEN subst]) haftmann@25600: apply (simp add: even_num_iff) haftmann@25600: done haftmann@25600: haftmann@25600: lemma lemma_even_mod_4_div_2: "n mod 4 = (1::nat) ==> even ((n - 1) div 2)" haftmann@25600: by (rule_tac t = n and n1 = 4 in mod_div_equality [THEN subst], simp) haftmann@25600: haftmann@25600: wenzelm@21263: text {* Simplify, when the exponent is a numeral *} wenzelm@21256: wenzelm@21256: lemmas power_0_left_number_of = power_0_left [of "number_of w", standard] wenzelm@21256: declare power_0_left_number_of [simp] wenzelm@21256: wenzelm@21263: lemmas zero_le_power_eq_number_of [simp] = wenzelm@21256: zero_le_power_eq [of _ "number_of w", standard] wenzelm@21256: wenzelm@21263: lemmas zero_less_power_eq_number_of [simp] = wenzelm@21256: zero_less_power_eq [of _ "number_of w", standard] wenzelm@21256: wenzelm@21263: lemmas power_le_zero_eq_number_of [simp] = wenzelm@21256: power_le_zero_eq [of _ "number_of w", standard] wenzelm@21256: wenzelm@21263: lemmas power_less_zero_eq_number_of [simp] = wenzelm@21256: power_less_zero_eq [of _ "number_of w", standard] wenzelm@21256: wenzelm@21263: lemmas zero_less_power_nat_eq_number_of [simp] = wenzelm@21256: zero_less_power_nat_eq [of _ "number_of w", standard] wenzelm@21256: wenzelm@21263: lemmas power_eq_0_iff_number_of [simp] = power_eq_0_iff [of _ "number_of w", standard] wenzelm@21256: wenzelm@21263: lemmas power_even_abs_number_of [simp] = power_even_abs [of "number_of w" _, standard] wenzelm@21256: wenzelm@21256: wenzelm@21256: subsection {* An Equivalence for @{term [source] "0 \ a^n"} *} wenzelm@21256: wenzelm@21256: lemma even_power_le_0_imp_0: wenzelm@21263: "a ^ (2*k) \ (0::'a::{ordered_idom,recpower}) ==> a=0" wenzelm@21263: by (induct k) (auto simp add: zero_le_mult_iff mult_le_0_iff power_Suc) wenzelm@21256: chaieb@23522: lemma zero_le_power_iff[presburger]: wenzelm@21263: "(0 \ a^n) = (0 \ (a::'a::{ordered_idom,recpower}) | even n)" wenzelm@21256: proof cases wenzelm@21256: assume even: "even n" wenzelm@21256: then obtain k where "n = 2*k" wenzelm@21256: by (auto simp add: even_nat_equiv_def2 numeral_2_eq_2) wenzelm@21263: thus ?thesis by (simp add: zero_le_even_power even) wenzelm@21256: next wenzelm@21256: assume odd: "odd n" wenzelm@21256: then obtain k where "n = Suc(2*k)" wenzelm@21256: by (auto simp add: odd_nat_equiv_def2 numeral_2_eq_2) wenzelm@21256: thus ?thesis wenzelm@21263: by (auto simp add: power_Suc zero_le_mult_iff zero_le_even_power wenzelm@21263: dest!: even_power_le_0_imp_0) wenzelm@21263: qed wenzelm@21263: wenzelm@21256: wenzelm@21256: subsection {* Miscellaneous *} wenzelm@21256: haftmann@25600: lemma odd_pos: "odd (n::nat) \ 0 < n" haftmann@25600: by (cases n, simp_all) haftmann@25600: chaieb@23522: lemma [presburger]:"(x + 1) div 2 = x div 2 \ even (x::int)" by presburger chaieb@23522: lemma [presburger]: "(x + 1) div 2 = x div 2 + 1 \ odd (x::int)" by presburger chaieb@23522: lemma even_plus_one_div_two: "even (x::int) ==> (x + 1) div 2 = x div 2" by presburger chaieb@23522: lemma odd_plus_one_div_two: "odd (x::int) ==> (x + 1) div 2 = x div 2 + 1" by presburger wenzelm@21256: chaieb@23522: lemma [presburger]: "(Suc x) div Suc (Suc 0) = x div Suc (Suc 0) \ even x" by presburger chaieb@23522: lemma [presburger]: "(Suc x) div Suc (Suc 0) = x div Suc (Suc 0) \ even x" by presburger wenzelm@21263: lemma even_nat_plus_one_div_two: "even (x::nat) ==> chaieb@23522: (Suc x) div Suc (Suc 0) = x div Suc (Suc 0)" by presburger wenzelm@21256: wenzelm@21263: lemma odd_nat_plus_one_div_two: "odd (x::nat) ==> chaieb@23522: (Suc x) div Suc (Suc 0) = Suc (x div Suc (Suc 0))" by presburger wenzelm@21256: wenzelm@21256: end