# HG changeset patch # User haftmann # Date 1555444205 0 # Node ID 56727602d0a5e7989066f37a80de69bc9975b7a0 # Parent 8bb835f10a393ae601a0bc75eaa7be5b9fa992e4 prefer one theory for misc material diff -r 8bb835f10a39 -r 56727602d0a5 src/HOL/Word/Bit_Representation.thy --- a/src/HOL/Word/Bit_Representation.thy Tue Apr 16 19:50:03 2019 +0000 +++ b/src/HOL/Word/Bit_Representation.thy Tue Apr 16 19:50:05 2019 +0000 @@ -5,9 +5,37 @@ section \Integers as implicit bit strings\ theory Bit_Representation - imports Misc_Numeric + imports Main begin +lemma int_mod_lem: "0 < n \ 0 \ b \ b < n \ b mod n = b" + for b n :: int + apply safe + apply (erule (1) mod_pos_pos_trivial) + apply (erule_tac [!] subst) + apply auto + done + +lemma int_mod_ge: "a < n \ 0 < n \ a \ a mod n" + for a n :: int + by (metis dual_order.trans le_cases mod_pos_pos_trivial pos_mod_conj) + +lemma int_mod_ge': "b < 0 \ 0 < n \ b + n \ b mod n" + for b n :: int + by (metis add_less_same_cancel2 int_mod_ge mod_add_self2) + +lemma int_mod_le': "0 \ b - n \ b mod n \ b - n" + for b n :: int + by (metis minus_mod_self2 zmod_le_nonneg_dividend) + +lemma emep1: "even n \ even d \ 0 \ d \ (n + 1) mod d = (n mod d) + 1" + for n d :: int + by (auto simp add: pos_zmod_mult_2 add.commute dvd_def) + +lemma m1mod2k: "- 1 mod 2 ^ n = (2 ^ n - 1 :: int)" + by (rule zmod_minus1) simp + + subsection \Constructors and destructors for binary integers\ definition Bit :: "int \ bool \ int" (infixl "BIT" 90) @@ -563,7 +591,9 @@ apply (unfold no_bintr_alt1) apply (auto simp add: image_iff) apply (rule exI) - apply (auto intro: int_mod_lem [THEN iffD1, symmetric]) + apply (rule sym) + using int_mod_lem [symmetric, of "2 ^ n"] + apply auto done lemma no_sbintr_alt2: "sbintrunc n = (\w. (w + 2 ^ n) mod 2 ^ Suc n - 2 ^ n :: int)" @@ -572,15 +602,15 @@ lemma range_sbintrunc: "range (sbintrunc n) = {i. - (2 ^ n) \ i \ i < 2 ^ n}" apply (unfold no_sbintr_alt2) apply (auto simp add: image_iff eq_diff_eq) + apply (rule exI) apply (auto intro: int_mod_lem [THEN iffD1, symmetric]) done lemma sb_inc_lem: "a + 2^k < 0 \ a + 2^k + 2^(Suc k) \ (a + 2^k) mod 2^(Suc k)" for a :: int - apply (erule int_mod_ge' [where n = "2 ^ (Suc k)" and b = "a + 2 ^ k", simplified zless2p]) - apply (rule TrueI) - done + using int_mod_ge' [where n = "2 ^ (Suc k)" and b = "a + 2 ^ k"] + by simp lemma sb_inc_lem': "a < - (2^k) \ a + 2^k + 2^(Suc k) \ (a + 2^k) mod 2^(Suc k)" for a :: int @@ -600,8 +630,6 @@ lemma sbintrunc_dec: "x \ (2 ^ n) \ x - 2 ^ (Suc n) >= sbintrunc n x" unfolding no_sbintr_alt2 by (drule sb_dec_lem') simp -lemmas m2pths = pos_mod_sign pos_mod_bound [OF zless2p] - lemma bintr_ge0: "0 \ bintrunc n w" by (simp add: bintrunc_mod2p) @@ -815,7 +843,7 @@ apply (simp add: bin_rest_def zdiv_zmult2_eq) apply (case_tac b rule: bin_exhaust) apply simp - apply (simp add: Bit_def mod_mult_mult1 p1mod22k) + apply (simp add: Bit_def mod_mult_mult1 pos_zmod_mult_2 add.commute) done end diff -r 8bb835f10a39 -r 56727602d0a5 src/HOL/Word/Bits_Int.thy --- a/src/HOL/Word/Bits_Int.thy Tue Apr 16 19:50:03 2019 +0000 +++ b/src/HOL/Word/Bits_Int.thy Tue Apr 16 19:50:05 2019 +0000 @@ -370,7 +370,7 @@ have "2 * (bin mod 2 ^ n) + 1 = (2 * bin mod 2 ^ Suc n) + 1" by (simp add: mod_mult_mult1) also have "\ = ((2 * bin mod 2 ^ Suc n) + 1) mod 2 ^ Suc n" - by (simp add: ac_simps p1mod22k') + by (simp add: ac_simps pos_zmod_mult_2) also have "\ = (2 * bin + 1) mod 2 ^ Suc n" by (simp only: mod_simps) finally show ?thesis diff -r 8bb835f10a39 -r 56727602d0a5 src/HOL/Word/Misc_Arithmetic.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Word/Misc_Arithmetic.thy Tue Apr 16 19:50:05 2019 +0000 @@ -0,0 +1,433 @@ +(* Title: HOL/Word/Misc_Arithmetic.thy *) + +section \Miscellaneous lemmas, mostly for arithmetic\ + +theory Misc_Arithmetic + imports "HOL-Library.Bit" Bit_Representation +begin + +lemma one_mod_exp_eq_one [simp]: + "1 mod (2 * 2 ^ n) = (1::int)" + using power_gt1 [of 2 n] by (auto intro: mod_pos_pos_trivial) + +lemma mod_2_neq_1_eq_eq_0: "k mod 2 \ 1 \ k mod 2 = 0" + for k :: int + by (fact not_mod_2_eq_1_eq_0) + +lemma z1pmod2: "(2 * b + 1) mod 2 = (1::int)" + for b :: int + by arith + +lemma diff_le_eq': "a - b \ c \ a \ b + c" + for a b c :: int + by arith + +lemma zless2: "0 < (2 :: int)" + by (fact zero_less_numeral) + +lemma zless2p: "0 < (2 ^ n :: int)" + by arith + +lemma zle2p: "0 \ (2 ^ n :: int)" + by arith + +lemma p1mod22k': "(1 + 2 * b) mod (2 * 2 ^ n) = 1 + 2 * (b mod 2 ^ n)" + for b :: int + using zle2p by (rule pos_zmod_mult_2) + +lemma p1mod22k: "(2 * b + 1) mod (2 * 2 ^ n) = 2 * (b mod 2 ^ n) + 1" + for b :: int + by (simp add: p1mod22k' add.commute) + +lemma ex_eq_or: "(\m. n = Suc m \ (m = k \ P m)) \ n = Suc k \ (\m. n = Suc m \ P m)" + by auto + +lemma power_minus_simp: "0 < n \ a ^ n = a * a ^ (n - 1)" + by (auto dest: gr0_implies_Suc) + +lemma funpow_minus_simp: "0 < n \ f ^^ n = f \ f ^^ (n - 1)" + by (auto dest: gr0_implies_Suc) + +lemma power_numeral: "a ^ numeral k = a * a ^ (pred_numeral k)" + by (simp add: numeral_eq_Suc) + +lemma funpow_numeral [simp]: "f ^^ numeral k = f \ f ^^ (pred_numeral k)" + by (simp add: numeral_eq_Suc) + +lemma replicate_numeral [simp]: "replicate (numeral k) x = x # replicate (pred_numeral k) x" + by (simp add: numeral_eq_Suc) + +lemma rco_alt: "(f \ g) ^^ n \ f = f \ (g \ f) ^^ n" + apply (rule ext) + apply (induct n) + apply (simp_all add: o_def) + done + +lemma list_exhaust_size_gt0: + assumes "\a list. y = a # list \ P" + shows "0 < length y \ P" + apply (cases y) + apply simp + apply (rule assms) + apply fastforce + done + +lemma list_exhaust_size_eq0: + assumes "y = [] \ P" + shows "length y = 0 \ P" + apply (cases y) + apply (rule assms) + apply simp + apply simp + done + +lemma size_Cons_lem_eq: "y = xa # list \ size y = Suc k \ size list = k" + by auto + +lemmas ls_splits = prod.split prod.split_asm if_split_asm + +\ \simplifications for specific word lengths\ +lemmas n2s_ths [THEN eq_reflection] = add_2_eq_Suc add_2_eq_Suc' + +lemmas s2n_ths = n2s_ths [symmetric] + +lemma and_len: "xs = ys \ xs = ys \ length xs = length ys" + by auto + +lemma size_if: "size (if p then xs else ys) = (if p then size xs else size ys)" + by auto + +lemma tl_if: "tl (if p then xs else ys) = (if p then tl xs else tl ys)" + by auto + +lemma hd_if: "hd (if p then xs else ys) = (if p then hd xs else hd ys)" + by auto + +lemma if_Not_x: "(if p then \ x else x) = (p = (\ x))" + by auto + +lemma if_x_Not: "(if p then x else \ x) = (p = x)" + by auto + +lemma if_same_and: "(If p x y \ If p u v) = (if p then x \ u else y \ v)" + by auto + +lemma if_same_eq: "(If p x y = (If p u v)) = (if p then x = u else y = v)" + by auto + +lemma if_same_eq_not: "(If p x y = (\ If p u v)) = (if p then x = (\ u) else y = (\ v))" + by auto + +\ \note -- \if_Cons\ can cause blowup in the size, if \p\ is complex, so make a simproc\ +lemma if_Cons: "(if p then x # xs else y # ys) = If p x y # If p xs ys" + by auto + +lemma if_single: "(if xc then [xab] else [an]) = [if xc then xab else an]" + by auto + +lemma if_bool_simps: + "If p True y = (p \ y) \ If p False y = (\ p \ y) \ + If p y True = (p \ y) \ If p y False = (p \ y)" + by auto + +lemmas if_simps = + if_x_Not if_Not_x if_cancel if_True if_False if_bool_simps + +lemmas seqr = eq_reflection [where x = "size w"] for w (* FIXME: delete *) + +lemma the_elemI: "y = {x} \ the_elem y = x" + by simp + +lemma nonemptyE: "S \ {} \ (\x. x \ S \ R) \ R" + by auto + +lemma gt_or_eq_0: "0 < y \ 0 = y" + for y :: nat + by arith + +lemmas xtr1 = xtrans(1) +lemmas xtr2 = xtrans(2) +lemmas xtr3 = xtrans(3) +lemmas xtr4 = xtrans(4) +lemmas xtr5 = xtrans(5) +lemmas xtr6 = xtrans(6) +lemmas xtr7 = xtrans(7) +lemmas xtr8 = xtrans(8) + +lemmas nat_simps = diff_add_inverse2 diff_add_inverse +lemmas nat_iffs = le_add1 le_add2 + +lemma sum_imp_diff: "j = k + i \ j - i = k" + for k :: nat + by arith + +lemmas pos_mod_sign2 = zless2 [THEN pos_mod_sign [where b = "2::int"]] +lemmas pos_mod_bound2 = zless2 [THEN pos_mod_bound [where b = "2::int"]] + +lemma nmod2: "n mod 2 = 0 \ n mod 2 = 1" + for n :: int + by arith + +lemmas eme1p = emep1 [simplified add.commute] + +lemma le_diff_eq': "a \ c - b \ b + a \ c" + for a b c :: int + by arith + +lemma less_diff_eq': "a < c - b \ b + a < c" + for a b c :: int + by arith + +lemma diff_less_eq': "a - b < c \ a < b + c" + for a b c :: int + by arith + +lemmas m1mod22k = mult_pos_pos [OF zless2 zless2p, THEN zmod_minus1] + +lemma z1pdiv2: "(2 * b + 1) div 2 = b" + for b :: int + by arith + +lemmas zdiv_le_dividend = xtr3 [OF div_by_1 [symmetric] zdiv_mono2, + simplified int_one_le_iff_zero_less, simplified] + +lemma axxbyy: "a + m + m = b + n + n \ a = 0 \ a = 1 \ b = 0 \ b = 1 \ a = b \ m = n" + for a b m n :: int + by arith + +lemma axxmod2: "(1 + x + x) mod 2 = 1 \ (0 + x + x) mod 2 = 0" + for x :: int + by arith + +lemma axxdiv2: "(1 + x + x) div 2 = x \ (0 + x + x) div 2 = x" + for x :: int + by arith + +lemmas iszero_minus = + trans [THEN trans, OF iszero_def neg_equal_0_iff_equal iszero_def [symmetric]] + +lemmas zadd_diff_inverse = + trans [OF diff_add_cancel [symmetric] add.commute] + +lemmas add_diff_cancel2 = + add.commute [THEN diff_eq_eq [THEN iffD2]] + +lemmas rdmods [symmetric] = mod_minus_eq + mod_diff_left_eq mod_diff_right_eq mod_add_left_eq + mod_add_right_eq mod_mult_right_eq mod_mult_left_eq + +lemma mod_plus_right: "(a + x) mod m = (b + x) mod m \ a mod m = b mod m" + for a b m x :: nat + by (induct x) (simp_all add: mod_Suc, arith) + +lemma nat_minus_mod: "(n - n mod m) mod m = 0" + for m n :: nat + by (induct n) (simp_all add: mod_Suc) + +lemmas nat_minus_mod_plus_right = + trans [OF nat_minus_mod mod_0 [symmetric], + THEN mod_plus_right [THEN iffD2], simplified] + +lemmas push_mods' = mod_add_eq + mod_mult_eq mod_diff_eq + mod_minus_eq + +lemmas push_mods = push_mods' [THEN eq_reflection] +lemmas pull_mods = push_mods [symmetric] rdmods [THEN eq_reflection] + +lemma nat_mod_eq: "b < n \ a mod n = b mod n \ a mod n = b" + for a b n :: nat + by (induct a) auto + +lemmas nat_mod_eq' = refl [THEN [2] nat_mod_eq] + +lemma nat_mod_lem: "0 < n \ b < n \ b mod n = b" + for b n :: nat + apply safe + apply (erule nat_mod_eq') + apply (erule subst) + apply (erule mod_less_divisor) + done + +lemma mod_nat_add: "x < z \ y < z \ (x + y) mod z = (if x + y < z then x + y else x + y - z)" + for x y z :: nat + apply (rule nat_mod_eq) + apply auto + apply (rule trans) + apply (rule le_mod_geq) + apply simp + apply (rule nat_mod_eq') + apply arith + done + +lemma mod_nat_sub: "x < z \ (x - y) mod z = x - y" + for x y :: nat + by (rule nat_mod_eq') arith + +lemma int_mod_eq: "0 \ b \ b < n \ a mod n = b mod n \ a mod n = b" + for a b n :: int + by (metis mod_pos_pos_trivial) + +lemmas int_mod_eq' = mod_pos_pos_trivial (* FIXME delete *) + +lemmas int_mod_le = zmod_le_nonneg_dividend (* FIXME: delete *) + +lemma mod_add_if_z: + "x < z \ y < z \ 0 \ y \ 0 \ x \ 0 \ z \ + (x + y) mod z = (if x + y < z then x + y else x + y - z)" + for x y z :: int + by (auto intro: int_mod_eq) + +lemma mod_sub_if_z: + "x < z \ y < z \ 0 \ y \ 0 \ x \ 0 \ z \ + (x - y) mod z = (if y \ x then x - y else x - y + z)" + for x y z :: int + by (auto intro: int_mod_eq) + +lemmas zmde = mult_div_mod_eq [symmetric, THEN diff_eq_eq [THEN iffD2], symmetric] +lemmas mcl = mult_cancel_left [THEN iffD1, THEN make_pos_rule] + +(* already have this for naturals, div_mult_self1/2, but not for ints *) +lemma zdiv_mult_self: "m \ 0 \ (a + m * n) div m = a div m + n" + for a m n :: int + apply (rule mcl) + prefer 2 + apply (erule asm_rl) + apply (simp add: zmde ring_distribs) + done + +lemma mod_power_lem: "a > 1 \ a ^ n mod a ^ m = (if m \ n then 0 else a ^ n)" + for a :: int + by (simp add: mod_eq_0_iff_dvd le_imp_power_dvd) + +lemma pl_pl_rels: "a + b = c + d \ a \ c \ b \ d \ a \ c \ b \ d" + for a b c d :: nat + by arith + +lemmas pl_pl_rels' = add.commute [THEN [2] trans, THEN pl_pl_rels] + +lemma minus_eq: "m - k = m \ k = 0 \ m = 0" + for k m :: nat + by arith + +lemma pl_pl_mm: "a + b = c + d \ a - c = d - b" + for a b c d :: nat + by arith + +lemmas pl_pl_mm' = add.commute [THEN [2] trans, THEN pl_pl_mm] + +lemmas dme = div_mult_mod_eq +lemmas dtle = div_times_less_eq_dividend +lemmas th2 = order_trans [OF order_refl [THEN [2] mult_le_mono] div_times_less_eq_dividend] + +lemma td_gal: "0 < c \ a \ b * c \ a div c \ b" + for a b c :: nat + apply safe + apply (erule (1) xtr4 [OF div_le_mono div_mult_self_is_m]) + apply (erule th2) + done + +lemmas td_gal_lt = td_gal [simplified not_less [symmetric], simplified] + +lemmas div_mult_le = div_times_less_eq_dividend + +lemmas sdl = div_nat_eqI + +lemma given_quot: "f > 0 \ (f * l + (f - 1)) div f = l" + for f l :: nat + by (rule div_nat_eqI) (simp_all) + +lemma given_quot_alt: "f > 0 \ (l * f + f - Suc 0) div f = l" + for f l :: nat + apply (frule given_quot) + apply (rule trans) + prefer 2 + apply (erule asm_rl) + apply (rule_tac f="\n. n div f" in arg_cong) + apply (simp add : ac_simps) + done + +lemma diff_mod_le: "a < d \ b dvd d \ a - a mod b \ d - b" + for a b d :: nat + apply (unfold dvd_def) + apply clarify + apply (case_tac k) + apply clarsimp + apply clarify + apply (cases "b > 0") + apply (drule mult.commute [THEN xtr1]) + apply (frule (1) td_gal_lt [THEN iffD1]) + apply (clarsimp simp: le_simps) + apply (rule minus_mod_eq_mult_div [symmetric, THEN [2] xtr4]) + apply (rule mult_mono) + apply auto + done + +lemma less_le_mult': "w * c < b * c \ 0 \ c \ (w + 1) * c \ b * c" + for b c w :: int + apply (rule mult_right_mono) + apply (rule zless_imp_add1_zle) + apply (erule (1) mult_right_less_imp_less) + apply assumption + done + +lemma less_le_mult: "w * c < b * c \ 0 \ c \ w * c + c \ b * c" + for b c w :: int + using less_le_mult' [of w c b] by (simp add: algebra_simps) + +lemmas less_le_mult_minus = iffD2 [OF le_diff_eq less_le_mult, + simplified left_diff_distrib] + +lemma gen_minus: "0 < n \ f n = f (Suc (n - 1))" + by auto + +lemma mpl_lem: "j \ i \ k < j \ i - j + k < i" + for i j k :: nat + by arith + +lemma nonneg_mod_div: "0 \ a \ 0 \ b \ 0 \ (a mod b) \ 0 \ a div b" + for a b :: int + by (cases "b = 0") (auto intro: pos_imp_zdiv_nonneg_iff [THEN iffD2]) + +declare iszero_0 [intro] + +lemma min_pm [simp]: "min a b + (a - b) = a" + for a b :: nat + by arith + +lemma min_pm1 [simp]: "a - b + min a b = a" + for a b :: nat + by arith + +lemma rev_min_pm [simp]: "min b a + (a - b) = a" + for a b :: nat + by arith + +lemma rev_min_pm1 [simp]: "a - b + min b a = a" + for a b :: nat + by arith + +lemma min_minus [simp]: "min m (m - k) = m - k" + for m k :: nat + by arith + +lemma min_minus' [simp]: "min (m - k) m = m - k" + for m k :: nat + by arith + +lemmas m2pths = pos_mod_sign pos_mod_bound [OF zless2p] + +lemma not_B1_is_B0: "y \ 1 \ y = 0" + for y :: bit + by (cases y) auto + +lemma B1_ass_B0: + fixes y :: bit + assumes y: "y = 0 \ y = 1" + shows "y = 1" + apply (rule classical) + apply (drule not_B1_is_B0) + apply (erule y) + done + +end diff -r 8bb835f10a39 -r 56727602d0a5 src/HOL/Word/Misc_Numeric.thy --- a/src/HOL/Word/Misc_Numeric.thy Tue Apr 16 19:50:03 2019 +0000 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,71 +0,0 @@ -(* Title: HOL/Word/Misc_Numeric.thy - Author: Jeremy Dawson, NICTA -*) - -section \Useful Numerical Lemmas\ - -theory Misc_Numeric - imports Main -begin - -lemma one_mod_exp_eq_one [simp]: - "1 mod (2 * 2 ^ n) = (1::int)" - using power_gt1 [of 2 n] by (auto intro: mod_pos_pos_trivial) - -lemma mod_2_neq_1_eq_eq_0: "k mod 2 \ 1 \ k mod 2 = 0" - for k :: int - by (fact not_mod_2_eq_1_eq_0) - -lemma z1pmod2: "(2 * b + 1) mod 2 = (1::int)" - for b :: int - by arith - -lemma diff_le_eq': "a - b \ c \ a \ b + c" - for a b c :: int - by arith - -lemma emep1: "even n \ even d \ 0 \ d \ (n + 1) mod d = (n mod d) + 1" - for n d :: int - by (auto simp add: pos_zmod_mult_2 add.commute dvd_def) - -lemma int_mod_ge: "a < n \ 0 < n \ a \ a mod n" - for a n :: int - by (metis dual_order.trans le_cases mod_pos_pos_trivial pos_mod_conj) - -lemma int_mod_ge': "b < 0 \ 0 < n \ b + n \ b mod n" - for b n :: int - by (metis add_less_same_cancel2 int_mod_ge mod_add_self2) - -lemma int_mod_le': "0 \ b - n \ b mod n \ b - n" - for b n :: int - by (metis minus_mod_self2 zmod_le_nonneg_dividend) - -lemma zless2: "0 < (2 :: int)" - by (fact zero_less_numeral) - -lemma zless2p: "0 < (2 ^ n :: int)" - by arith - -lemma zle2p: "0 \ (2 ^ n :: int)" - by arith - -lemma m1mod2k: "- 1 mod 2 ^ n = (2 ^ n - 1 :: int)" - using zless2p by (rule zmod_minus1) - -lemma p1mod22k': "(1 + 2 * b) mod (2 * 2 ^ n) = 1 + 2 * (b mod 2 ^ n)" - for b :: int - using zle2p by (rule pos_zmod_mult_2) - -lemma p1mod22k: "(2 * b + 1) mod (2 * 2 ^ n) = 2 * (b mod 2 ^ n) + 1" - for b :: int - by (simp add: p1mod22k' add.commute) - -lemma int_mod_lem: "0 < n \ 0 \ b \ b < n \ b mod n = b" - for b n :: int - apply safe - apply (erule (1) mod_pos_pos_trivial) - apply (erule_tac [!] subst) - apply auto - done - -end diff -r 8bb835f10a39 -r 56727602d0a5 src/HOL/Word/Word.thy --- a/src/HOL/Word/Word.thy Tue Apr 16 19:50:03 2019 +0000 +++ b/src/HOL/Word/Word.thy Tue Apr 16 19:50:05 2019 +0000 @@ -11,7 +11,7 @@ Bits_Bit Bits_Int Misc_Typedef - Word_Miscellaneous + Misc_Arithmetic begin text \See \<^file>\WordExamples.thy\ for examples.\ diff -r 8bb835f10a39 -r 56727602d0a5 src/HOL/Word/Word_Miscellaneous.thy --- a/src/HOL/Word/Word_Miscellaneous.thy Tue Apr 16 19:50:03 2019 +0000 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,398 +0,0 @@ -(* Title: HOL/Word/Word_Miscellaneous.thy *) - -section \Miscellaneous lemmas, of at least doubtful value\ - -theory Word_Miscellaneous - imports "HOL-Library.Bit" Misc_Numeric -begin - -lemma ex_eq_or: "(\m. n = Suc m \ (m = k \ P m)) \ n = Suc k \ (\m. n = Suc m \ P m)" - by auto - -lemma power_minus_simp: "0 < n \ a ^ n = a * a ^ (n - 1)" - by (auto dest: gr0_implies_Suc) - -lemma funpow_minus_simp: "0 < n \ f ^^ n = f \ f ^^ (n - 1)" - by (auto dest: gr0_implies_Suc) - -lemma power_numeral: "a ^ numeral k = a * a ^ (pred_numeral k)" - by (simp add: numeral_eq_Suc) - -lemma funpow_numeral [simp]: "f ^^ numeral k = f \ f ^^ (pred_numeral k)" - by (simp add: numeral_eq_Suc) - -lemma replicate_numeral [simp]: "replicate (numeral k) x = x # replicate (pred_numeral k) x" - by (simp add: numeral_eq_Suc) - -lemma rco_alt: "(f \ g) ^^ n \ f = f \ (g \ f) ^^ n" - apply (rule ext) - apply (induct n) - apply (simp_all add: o_def) - done - -lemma list_exhaust_size_gt0: - assumes "\a list. y = a # list \ P" - shows "0 < length y \ P" - apply (cases y) - apply simp - apply (rule assms) - apply fastforce - done - -lemma list_exhaust_size_eq0: - assumes "y = [] \ P" - shows "length y = 0 \ P" - apply (cases y) - apply (rule assms) - apply simp - apply simp - done - -lemma size_Cons_lem_eq: "y = xa # list \ size y = Suc k \ size list = k" - by auto - -lemmas ls_splits = prod.split prod.split_asm if_split_asm - -lemma not_B1_is_B0: "y \ 1 \ y = 0" - for y :: bit - by (cases y) auto - -lemma B1_ass_B0: - fixes y :: bit - assumes y: "y = 0 \ y = 1" - shows "y = 1" - apply (rule classical) - apply (drule not_B1_is_B0) - apply (erule y) - done - -\ \simplifications for specific word lengths\ -lemmas n2s_ths [THEN eq_reflection] = add_2_eq_Suc add_2_eq_Suc' - -lemmas s2n_ths = n2s_ths [symmetric] - -lemma and_len: "xs = ys \ xs = ys \ length xs = length ys" - by auto - -lemma size_if: "size (if p then xs else ys) = (if p then size xs else size ys)" - by auto - -lemma tl_if: "tl (if p then xs else ys) = (if p then tl xs else tl ys)" - by auto - -lemma hd_if: "hd (if p then xs else ys) = (if p then hd xs else hd ys)" - by auto - -lemma if_Not_x: "(if p then \ x else x) = (p = (\ x))" - by auto - -lemma if_x_Not: "(if p then x else \ x) = (p = x)" - by auto - -lemma if_same_and: "(If p x y \ If p u v) = (if p then x \ u else y \ v)" - by auto - -lemma if_same_eq: "(If p x y = (If p u v)) = (if p then x = u else y = v)" - by auto - -lemma if_same_eq_not: "(If p x y = (\ If p u v)) = (if p then x = (\ u) else y = (\ v))" - by auto - -\ \note -- \if_Cons\ can cause blowup in the size, if \p\ is complex, so make a simproc\ -lemma if_Cons: "(if p then x # xs else y # ys) = If p x y # If p xs ys" - by auto - -lemma if_single: "(if xc then [xab] else [an]) = [if xc then xab else an]" - by auto - -lemma if_bool_simps: - "If p True y = (p \ y) \ If p False y = (\ p \ y) \ - If p y True = (p \ y) \ If p y False = (p \ y)" - by auto - -lemmas if_simps = - if_x_Not if_Not_x if_cancel if_True if_False if_bool_simps - -lemmas seqr = eq_reflection [where x = "size w"] for w (* FIXME: delete *) - -lemma the_elemI: "y = {x} \ the_elem y = x" - by simp - -lemma nonemptyE: "S \ {} \ (\x. x \ S \ R) \ R" - by auto - -lemma gt_or_eq_0: "0 < y \ 0 = y" - for y :: nat - by arith - -lemmas xtr1 = xtrans(1) -lemmas xtr2 = xtrans(2) -lemmas xtr3 = xtrans(3) -lemmas xtr4 = xtrans(4) -lemmas xtr5 = xtrans(5) -lemmas xtr6 = xtrans(6) -lemmas xtr7 = xtrans(7) -lemmas xtr8 = xtrans(8) - -lemmas nat_simps = diff_add_inverse2 diff_add_inverse -lemmas nat_iffs = le_add1 le_add2 - -lemma sum_imp_diff: "j = k + i \ j - i = k" - for k :: nat - by arith - -lemmas pos_mod_sign2 = zless2 [THEN pos_mod_sign [where b = "2::int"]] -lemmas pos_mod_bound2 = zless2 [THEN pos_mod_bound [where b = "2::int"]] - -lemma nmod2: "n mod 2 = 0 \ n mod 2 = 1" - for n :: int - by arith - -lemmas eme1p = emep1 [simplified add.commute] - -lemma le_diff_eq': "a \ c - b \ b + a \ c" - for a b c :: int - by arith - -lemma less_diff_eq': "a < c - b \ b + a < c" - for a b c :: int - by arith - -lemma diff_less_eq': "a - b < c \ a < b + c" - for a b c :: int - by arith - -lemmas m1mod22k = mult_pos_pos [OF zless2 zless2p, THEN zmod_minus1] - -lemma z1pdiv2: "(2 * b + 1) div 2 = b" - for b :: int - by arith - -lemmas zdiv_le_dividend = xtr3 [OF div_by_1 [symmetric] zdiv_mono2, - simplified int_one_le_iff_zero_less, simplified] - -lemma axxbyy: "a + m + m = b + n + n \ a = 0 \ a = 1 \ b = 0 \ b = 1 \ a = b \ m = n" - for a b m n :: int - by arith - -lemma axxmod2: "(1 + x + x) mod 2 = 1 \ (0 + x + x) mod 2 = 0" - for x :: int - by arith - -lemma axxdiv2: "(1 + x + x) div 2 = x \ (0 + x + x) div 2 = x" - for x :: int - by arith - -lemmas iszero_minus = - trans [THEN trans, OF iszero_def neg_equal_0_iff_equal iszero_def [symmetric]] - -lemmas zadd_diff_inverse = - trans [OF diff_add_cancel [symmetric] add.commute] - -lemmas add_diff_cancel2 = - add.commute [THEN diff_eq_eq [THEN iffD2]] - -lemmas rdmods [symmetric] = mod_minus_eq - mod_diff_left_eq mod_diff_right_eq mod_add_left_eq - mod_add_right_eq mod_mult_right_eq mod_mult_left_eq - -lemma mod_plus_right: "(a + x) mod m = (b + x) mod m \ a mod m = b mod m" - for a b m x :: nat - by (induct x) (simp_all add: mod_Suc, arith) - -lemma nat_minus_mod: "(n - n mod m) mod m = 0" - for m n :: nat - by (induct n) (simp_all add: mod_Suc) - -lemmas nat_minus_mod_plus_right = - trans [OF nat_minus_mod mod_0 [symmetric], - THEN mod_plus_right [THEN iffD2], simplified] - -lemmas push_mods' = mod_add_eq - mod_mult_eq mod_diff_eq - mod_minus_eq - -lemmas push_mods = push_mods' [THEN eq_reflection] -lemmas pull_mods = push_mods [symmetric] rdmods [THEN eq_reflection] - -lemma nat_mod_eq: "b < n \ a mod n = b mod n \ a mod n = b" - for a b n :: nat - by (induct a) auto - -lemmas nat_mod_eq' = refl [THEN [2] nat_mod_eq] - -lemma nat_mod_lem: "0 < n \ b < n \ b mod n = b" - for b n :: nat - apply safe - apply (erule nat_mod_eq') - apply (erule subst) - apply (erule mod_less_divisor) - done - -lemma mod_nat_add: "x < z \ y < z \ (x + y) mod z = (if x + y < z then x + y else x + y - z)" - for x y z :: nat - apply (rule nat_mod_eq) - apply auto - apply (rule trans) - apply (rule le_mod_geq) - apply simp - apply (rule nat_mod_eq') - apply arith - done - -lemma mod_nat_sub: "x < z \ (x - y) mod z = x - y" - for x y :: nat - by (rule nat_mod_eq') arith - -lemma int_mod_eq: "0 \ b \ b < n \ a mod n = b mod n \ a mod n = b" - for a b n :: int - by (metis mod_pos_pos_trivial) - -lemmas int_mod_eq' = mod_pos_pos_trivial (* FIXME delete *) - -lemmas int_mod_le = zmod_le_nonneg_dividend (* FIXME: delete *) - -lemma mod_add_if_z: - "x < z \ y < z \ 0 \ y \ 0 \ x \ 0 \ z \ - (x + y) mod z = (if x + y < z then x + y else x + y - z)" - for x y z :: int - by (auto intro: int_mod_eq) - -lemma mod_sub_if_z: - "x < z \ y < z \ 0 \ y \ 0 \ x \ 0 \ z \ - (x - y) mod z = (if y \ x then x - y else x - y + z)" - for x y z :: int - by (auto intro: int_mod_eq) - -lemmas zmde = mult_div_mod_eq [symmetric, THEN diff_eq_eq [THEN iffD2], symmetric] -lemmas mcl = mult_cancel_left [THEN iffD1, THEN make_pos_rule] - -(* already have this for naturals, div_mult_self1/2, but not for ints *) -lemma zdiv_mult_self: "m \ 0 \ (a + m * n) div m = a div m + n" - for a m n :: int - apply (rule mcl) - prefer 2 - apply (erule asm_rl) - apply (simp add: zmde ring_distribs) - done - -lemma mod_power_lem: "a > 1 \ a ^ n mod a ^ m = (if m \ n then 0 else a ^ n)" - for a :: int - by (simp add: mod_eq_0_iff_dvd le_imp_power_dvd) - -lemma pl_pl_rels: "a + b = c + d \ a \ c \ b \ d \ a \ c \ b \ d" - for a b c d :: nat - by arith - -lemmas pl_pl_rels' = add.commute [THEN [2] trans, THEN pl_pl_rels] - -lemma minus_eq: "m - k = m \ k = 0 \ m = 0" - for k m :: nat - by arith - -lemma pl_pl_mm: "a + b = c + d \ a - c = d - b" - for a b c d :: nat - by arith - -lemmas pl_pl_mm' = add.commute [THEN [2] trans, THEN pl_pl_mm] - -lemmas dme = div_mult_mod_eq -lemmas dtle = div_times_less_eq_dividend -lemmas th2 = order_trans [OF order_refl [THEN [2] mult_le_mono] div_times_less_eq_dividend] - -lemma td_gal: "0 < c \ a \ b * c \ a div c \ b" - for a b c :: nat - apply safe - apply (erule (1) xtr4 [OF div_le_mono div_mult_self_is_m]) - apply (erule th2) - done - -lemmas td_gal_lt = td_gal [simplified not_less [symmetric], simplified] - -lemmas div_mult_le = div_times_less_eq_dividend - -lemmas sdl = div_nat_eqI - -lemma given_quot: "f > 0 \ (f * l + (f - 1)) div f = l" - for f l :: nat - by (rule div_nat_eqI) (simp_all) - -lemma given_quot_alt: "f > 0 \ (l * f + f - Suc 0) div f = l" - for f l :: nat - apply (frule given_quot) - apply (rule trans) - prefer 2 - apply (erule asm_rl) - apply (rule_tac f="\n. n div f" in arg_cong) - apply (simp add : ac_simps) - done - -lemma diff_mod_le: "a < d \ b dvd d \ a - a mod b \ d - b" - for a b d :: nat - apply (unfold dvd_def) - apply clarify - apply (case_tac k) - apply clarsimp - apply clarify - apply (cases "b > 0") - apply (drule mult.commute [THEN xtr1]) - apply (frule (1) td_gal_lt [THEN iffD1]) - apply (clarsimp simp: le_simps) - apply (rule minus_mod_eq_mult_div [symmetric, THEN [2] xtr4]) - apply (rule mult_mono) - apply auto - done - -lemma less_le_mult': "w * c < b * c \ 0 \ c \ (w + 1) * c \ b * c" - for b c w :: int - apply (rule mult_right_mono) - apply (rule zless_imp_add1_zle) - apply (erule (1) mult_right_less_imp_less) - apply assumption - done - -lemma less_le_mult: "w * c < b * c \ 0 \ c \ w * c + c \ b * c" - for b c w :: int - using less_le_mult' [of w c b] by (simp add: algebra_simps) - -lemmas less_le_mult_minus = iffD2 [OF le_diff_eq less_le_mult, - simplified left_diff_distrib] - -lemma gen_minus: "0 < n \ f n = f (Suc (n - 1))" - by auto - -lemma mpl_lem: "j \ i \ k < j \ i - j + k < i" - for i j k :: nat - by arith - -lemma nonneg_mod_div: "0 \ a \ 0 \ b \ 0 \ (a mod b) \ 0 \ a div b" - for a b :: int - by (cases "b = 0") (auto intro: pos_imp_zdiv_nonneg_iff [THEN iffD2]) - -declare iszero_0 [intro] - -lemma min_pm [simp]: "min a b + (a - b) = a" - for a b :: nat - by arith - -lemma min_pm1 [simp]: "a - b + min a b = a" - for a b :: nat - by arith - -lemma rev_min_pm [simp]: "min b a + (a - b) = a" - for a b :: nat - by arith - -lemma rev_min_pm1 [simp]: "a - b + min b a = a" - for a b :: nat - by arith - -lemma min_minus [simp]: "min m (m - k) = m - k" - for m k :: nat - by arith - -lemma min_minus' [simp]: "min (m - k) m = m - k" - for m k :: nat - by arith - -end