| author | wenzelm | 
| Sun, 11 Aug 2019 22:36:34 +0200 | |
| changeset 70503 | f0b2635ee17f | 
| parent 70342 | e4d626692640 | 
| child 71987 | ec17263ec38d | 
| permissions | -rw-r--r-- | 
| 70170 | 1 | (* Title: HOL/Word/Misc_Arithmetic.thy *) | 
| 2 | ||
| 3 | section \<open>Miscellaneous lemmas, mostly for arithmetic\<close> | |
| 4 | ||
| 5 | theory Misc_Arithmetic | |
| 70342 
e4d626692640
clear separation of types for bits (False / True) and Z2 (0 / 1)
 haftmann parents: 
70190diff
changeset | 6 | imports Misc_Auxiliary "HOL-Library.Z2" | 
| 70170 | 7 | begin | 
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 8 | |
| 70170 | 9 | lemma one_mod_exp_eq_one [simp]: | 
| 10 | "1 mod (2 * 2 ^ n) = (1::int)" | |
| 11 | using power_gt1 [of 2 n] by (auto intro: mod_pos_pos_trivial) | |
| 12 | ||
| 13 | lemma mod_2_neq_1_eq_eq_0: "k mod 2 \<noteq> 1 \<longleftrightarrow> k mod 2 = 0" | |
| 14 | for k :: int | |
| 15 | by (fact not_mod_2_eq_1_eq_0) | |
| 16 | ||
| 17 | lemma z1pmod2: "(2 * b + 1) mod 2 = (1::int)" | |
| 18 | for b :: int | |
| 19 | by arith | |
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 20 | |
| 70170 | 21 | lemma diff_le_eq': "a - b \<le> c \<longleftrightarrow> a \<le> b + c" | 
| 22 | for a b c :: int | |
| 23 | by arith | |
| 24 | ||
| 25 | lemma zless2: "0 < (2 :: int)" | |
| 26 | by (fact zero_less_numeral) | |
| 27 | ||
| 28 | lemma zless2p: "0 < (2 ^ n :: int)" | |
| 29 | by arith | |
| 30 | ||
| 31 | lemma zle2p: "0 \<le> (2 ^ n :: int)" | |
| 32 | by arith | |
| 33 | ||
| 34 | lemma p1mod22k': "(1 + 2 * b) mod (2 * 2 ^ n) = 1 + 2 * (b mod 2 ^ n)" | |
| 35 | for b :: int | |
| 36 | using zle2p by (rule pos_zmod_mult_2) | |
| 37 | ||
| 38 | lemma p1mod22k: "(2 * b + 1) mod (2 * 2 ^ n) = 2 * (b mod 2 ^ n) + 1" | |
| 39 | for b :: int | |
| 40 | by (simp add: p1mod22k' add.commute) | |
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 41 | |
| 70169 | 42 | lemma ex_eq_or: "(\<exists>m. n = Suc m \<and> (m = k \<or> P m)) \<longleftrightarrow> n = Suc k \<or> (\<exists>m. n = Suc m \<and> P m)" | 
| 43 | by auto | |
| 44 | ||
| 65363 | 45 | lemma power_minus_simp: "0 < n \<Longrightarrow> a ^ n = a * a ^ (n - 1)" | 
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 46 | by (auto dest: gr0_implies_Suc) | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 47 | |
| 65363 | 48 | lemma funpow_minus_simp: "0 < n \<Longrightarrow> f ^^ n = f \<circ> f ^^ (n - 1)" | 
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 49 | by (auto dest: gr0_implies_Suc) | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 50 | |
| 65363 | 51 | lemma power_numeral: "a ^ numeral k = a * a ^ (pred_numeral k)" | 
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 52 | by (simp add: numeral_eq_Suc) | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 53 | |
| 65363 | 54 | lemma funpow_numeral [simp]: "f ^^ numeral k = f \<circ> f ^^ (pred_numeral k)" | 
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 55 | by (simp add: numeral_eq_Suc) | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 56 | |
| 65363 | 57 | lemma replicate_numeral [simp]: "replicate (numeral k) x = x # replicate (pred_numeral k) x" | 
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 58 | by (simp add: numeral_eq_Suc) | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 59 | |
| 65363 | 60 | lemma rco_alt: "(f \<circ> g) ^^ n \<circ> f = f \<circ> (g \<circ> f) ^^ n" | 
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 61 | apply (rule ext) | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 62 | apply (induct n) | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 63 | apply (simp_all add: o_def) | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 64 | done | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 65 | |
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 66 | lemma list_exhaust_size_gt0: | 
| 67120 | 67 | assumes "\<And>a list. y = a # list \<Longrightarrow> P" | 
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 68 | shows "0 < length y \<Longrightarrow> P" | 
| 65363 | 69 | apply (cases y) | 
| 70 | apply simp | |
| 67120 | 71 | apply (rule assms) | 
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 72 | apply fastforce | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 73 | done | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 74 | |
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 75 | lemma list_exhaust_size_eq0: | 
| 67120 | 76 | assumes "y = [] \<Longrightarrow> P" | 
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 77 | shows "length y = 0 \<Longrightarrow> P" | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 78 | apply (cases y) | 
| 67120 | 79 | apply (rule assms) | 
| 65363 | 80 | apply simp | 
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 81 | apply simp | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 82 | done | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 83 | |
| 65363 | 84 | lemma size_Cons_lem_eq: "y = xa # list \<Longrightarrow> size y = Suc k \<Longrightarrow> size list = k" | 
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 85 | by auto | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 86 | |
| 62390 | 87 | lemmas ls_splits = prod.split prod.split_asm if_split_asm | 
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 88 | |
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
67408diff
changeset | 89 | \<comment> \<open>simplifications for specific word lengths\<close> | 
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 90 | lemmas n2s_ths [THEN eq_reflection] = add_2_eq_Suc add_2_eq_Suc' | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 91 | |
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 92 | lemmas s2n_ths = n2s_ths [symmetric] | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 93 | |
| 65363 | 94 | lemma and_len: "xs = ys \<Longrightarrow> xs = ys \<and> length xs = length ys" | 
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 95 | by auto | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 96 | |
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 97 | lemma size_if: "size (if p then xs else ys) = (if p then size xs else size ys)" | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 98 | by auto | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 99 | |
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 100 | lemma tl_if: "tl (if p then xs else ys) = (if p then tl xs else tl ys)" | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 101 | by auto | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 102 | |
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 103 | lemma hd_if: "hd (if p then xs else ys) = (if p then hd xs else hd ys)" | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 104 | by auto | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 105 | |
| 65363 | 106 | lemma if_Not_x: "(if p then \<not> x else x) = (p = (\<not> x))" | 
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 107 | by auto | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 108 | |
| 65363 | 109 | lemma if_x_Not: "(if p then x else \<not> x) = (p = x)" | 
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 110 | by auto | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 111 | |
| 65363 | 112 | lemma if_same_and: "(If p x y \<and> If p u v) = (if p then x \<and> u else y \<and> v)" | 
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 113 | by auto | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 114 | |
| 65363 | 115 | lemma if_same_eq: "(If p x y = (If p u v)) = (if p then x = u else y = v)" | 
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 116 | by auto | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 117 | |
| 65363 | 118 | lemma if_same_eq_not: "(If p x y = (\<not> If p u v)) = (if p then x = (\<not> u) else y = (\<not> v))" | 
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 119 | by auto | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 120 | |
| 67408 | 121 | \<comment> \<open>note -- \<open>if_Cons\<close> can cause blowup in the size, if \<open>p\<close> is complex, so make a simproc\<close> | 
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 122 | lemma if_Cons: "(if p then x # xs else y # ys) = If p x y # If p xs ys" | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 123 | by auto | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 124 | |
| 65363 | 125 | lemma if_single: "(if xc then [xab] else [an]) = [if xc then xab else an]" | 
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 126 | by auto | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 127 | |
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 128 | lemma if_bool_simps: | 
| 65363 | 129 | "If p True y = (p \<or> y) \<and> If p False y = (\<not> p \<and> y) \<and> | 
| 130 | If p y True = (p \<longrightarrow> y) \<and> If p y False = (p \<and> y)" | |
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 131 | by auto | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 132 | |
| 65363 | 133 | lemmas if_simps = | 
| 134 | if_x_Not if_Not_x if_cancel if_True if_False if_bool_simps | |
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 135 | |
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 136 | lemmas seqr = eq_reflection [where x = "size w"] for w (* FIXME: delete *) | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 137 | |
| 65363 | 138 | lemma the_elemI: "y = {x} \<Longrightarrow> the_elem y = x"
 | 
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 139 | by simp | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 140 | |
| 65363 | 141 | lemma nonemptyE: "S \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> S \<Longrightarrow> R) \<Longrightarrow> R"
 | 
| 142 | by auto | |
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 143 | |
| 65363 | 144 | lemma gt_or_eq_0: "0 < y \<or> 0 = y" | 
| 145 | for y :: nat | |
| 146 | by arith | |
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 147 | |
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 148 | lemmas xtr1 = xtrans(1) | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 149 | lemmas xtr2 = xtrans(2) | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 150 | lemmas xtr3 = xtrans(3) | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 151 | lemmas xtr4 = xtrans(4) | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 152 | lemmas xtr5 = xtrans(5) | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 153 | lemmas xtr6 = xtrans(6) | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 154 | lemmas xtr7 = xtrans(7) | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 155 | lemmas xtr8 = xtrans(8) | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 156 | |
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 157 | lemmas nat_simps = diff_add_inverse2 diff_add_inverse | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 158 | lemmas nat_iffs = le_add1 le_add2 | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 159 | |
| 65363 | 160 | lemma sum_imp_diff: "j = k + i \<Longrightarrow> j - i = k" | 
| 161 | for k :: nat | |
| 162 | by arith | |
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 163 | |
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 164 | lemmas pos_mod_sign2 = zless2 [THEN pos_mod_sign [where b = "2::int"]] | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 165 | lemmas pos_mod_bound2 = zless2 [THEN pos_mod_bound [where b = "2::int"]] | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 166 | |
| 65363 | 167 | lemma nmod2: "n mod 2 = 0 \<or> n mod 2 = 1" | 
| 168 | for n :: int | |
| 58681 | 169 | by arith | 
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 170 | |
| 70190 | 171 | lemma eme1p: | 
| 172 | "even n \<Longrightarrow> even d \<Longrightarrow> 0 \<le> d \<Longrightarrow> (1 + n) mod d = 1 + n mod d" for n d :: int | |
| 173 | using emep1 [of n d] by (simp add: ac_simps) | |
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 174 | |
| 65363 | 175 | lemma le_diff_eq': "a \<le> c - b \<longleftrightarrow> b + a \<le> c" | 
| 176 | for a b c :: int | |
| 177 | by arith | |
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 178 | |
| 65363 | 179 | lemma less_diff_eq': "a < c - b \<longleftrightarrow> b + a < c" | 
| 180 | for a b c :: int | |
| 181 | by arith | |
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 182 | |
| 65363 | 183 | lemma diff_less_eq': "a - b < c \<longleftrightarrow> a < b + c" | 
| 184 | for a b c :: int | |
| 185 | by arith | |
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 186 | |
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 187 | lemmas m1mod22k = mult_pos_pos [OF zless2 zless2p, THEN zmod_minus1] | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 188 | |
| 65363 | 189 | lemma z1pdiv2: "(2 * b + 1) div 2 = b" | 
| 190 | for b :: int | |
| 191 | by arith | |
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 192 | |
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 193 | lemmas zdiv_le_dividend = xtr3 [OF div_by_1 [symmetric] zdiv_mono2, | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 194 | simplified int_one_le_iff_zero_less, simplified] | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 195 | |
| 65363 | 196 | lemma axxbyy: "a + m + m = b + n + n \<Longrightarrow> a = 0 \<or> a = 1 \<Longrightarrow> b = 0 \<or> b = 1 \<Longrightarrow> a = b \<and> m = n" | 
| 197 | for a b m n :: int | |
| 198 | by arith | |
| 199 | ||
| 200 | lemma axxmod2: "(1 + x + x) mod 2 = 1 \<and> (0 + x + x) mod 2 = 0" | |
| 201 | for x :: int | |
| 202 | by arith | |
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 203 | |
| 65363 | 204 | lemma axxdiv2: "(1 + x + x) div 2 = x \<and> (0 + x + x) div 2 = x" | 
| 205 | for x :: int | |
| 206 | by arith | |
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 207 | |
| 65363 | 208 | lemmas iszero_minus = | 
| 209 | trans [THEN trans, OF iszero_def neg_equal_0_iff_equal iszero_def [symmetric]] | |
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 210 | |
| 65363 | 211 | lemmas zadd_diff_inverse = | 
| 212 | trans [OF diff_add_cancel [symmetric] add.commute] | |
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 213 | |
| 65363 | 214 | lemmas add_diff_cancel2 = | 
| 215 | add.commute [THEN diff_eq_eq [THEN iffD2]] | |
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 216 | |
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 217 | lemmas rdmods [symmetric] = mod_minus_eq | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 218 | mod_diff_left_eq mod_diff_right_eq mod_add_left_eq | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 219 | mod_add_right_eq mod_mult_right_eq mod_mult_left_eq | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 220 | |
| 65363 | 221 | lemma mod_plus_right: "(a + x) mod m = (b + x) mod m \<longleftrightarrow> a mod m = b mod m" | 
| 222 | for a b m x :: nat | |
| 223 | by (induct x) (simp_all add: mod_Suc, arith) | |
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 224 | |
| 65363 | 225 | lemma nat_minus_mod: "(n - n mod m) mod m = 0" | 
| 226 | for m n :: nat | |
| 227 | by (induct n) (simp_all add: mod_Suc) | |
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 228 | |
| 65363 | 229 | lemmas nat_minus_mod_plus_right = | 
| 230 | trans [OF nat_minus_mod mod_0 [symmetric], | |
| 231 | THEN mod_plus_right [THEN iffD2], simplified] | |
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 232 | |
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 233 | lemmas push_mods' = mod_add_eq | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 234 | mod_mult_eq mod_diff_eq | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 235 | mod_minus_eq | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 236 | |
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 237 | lemmas push_mods = push_mods' [THEN eq_reflection] | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 238 | lemmas pull_mods = push_mods [symmetric] rdmods [THEN eq_reflection] | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 239 | |
| 65363 | 240 | lemma nat_mod_eq: "b < n \<Longrightarrow> a mod n = b mod n \<Longrightarrow> a mod n = b" | 
| 241 | for a b n :: nat | |
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 242 | by (induct a) auto | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 243 | |
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 244 | lemmas nat_mod_eq' = refl [THEN [2] nat_mod_eq] | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 245 | |
| 65363 | 246 | lemma nat_mod_lem: "0 < n \<Longrightarrow> b < n \<longleftrightarrow> b mod n = b" | 
| 247 | for b n :: nat | |
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 248 | apply safe | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 249 | apply (erule nat_mod_eq') | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 250 | apply (erule subst) | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 251 | apply (erule mod_less_divisor) | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 252 | done | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 253 | |
| 65363 | 254 | lemma mod_nat_add: "x < z \<Longrightarrow> y < z \<Longrightarrow> (x + y) mod z = (if x + y < z then x + y else x + y - z)" | 
| 255 | for x y z :: nat | |
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 256 | apply (rule nat_mod_eq) | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 257 | apply auto | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 258 | apply (rule trans) | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 259 | apply (rule le_mod_geq) | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 260 | apply simp | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 261 | apply (rule nat_mod_eq') | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 262 | apply arith | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 263 | done | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 264 | |
| 65363 | 265 | lemma mod_nat_sub: "x < z \<Longrightarrow> (x - y) mod z = x - y" | 
| 266 | for x y :: nat | |
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 267 | by (rule nat_mod_eq') arith | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 268 | |
| 65363 | 269 | lemma int_mod_eq: "0 \<le> b \<Longrightarrow> b < n \<Longrightarrow> a mod n = b mod n \<Longrightarrow> a mod n = b" | 
| 270 | for a b n :: int | |
| 55816 
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
 haftmann parents: 
55417diff
changeset | 271 | by (metis mod_pos_pos_trivial) | 
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 272 | |
| 55816 
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
 haftmann parents: 
55417diff
changeset | 273 | lemmas int_mod_eq' = mod_pos_pos_trivial (* FIXME delete *) | 
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 274 | |
| 66801 | 275 | lemmas int_mod_le = zmod_le_nonneg_dividend (* FIXME: delete *) | 
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 276 | |
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 277 | lemma mod_add_if_z: | 
| 65363 | 278 | "x < z \<Longrightarrow> y < z \<Longrightarrow> 0 \<le> y \<Longrightarrow> 0 \<le> x \<Longrightarrow> 0 \<le> z \<Longrightarrow> | 
| 279 | (x + y) mod z = (if x + y < z then x + y else x + y - z)" | |
| 280 | for x y z :: int | |
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 281 | by (auto intro: int_mod_eq) | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 282 | |
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 283 | lemma mod_sub_if_z: | 
| 65363 | 284 | "x < z \<Longrightarrow> y < z \<Longrightarrow> 0 \<le> y \<Longrightarrow> 0 \<le> x \<Longrightarrow> 0 \<le> z \<Longrightarrow> | 
| 285 | (x - y) mod z = (if y \<le> x then x - y else x - y + z)" | |
| 286 | for x y z :: int | |
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 287 | by (auto intro: int_mod_eq) | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 288 | |
| 64246 | 289 | lemmas zmde = mult_div_mod_eq [symmetric, THEN diff_eq_eq [THEN iffD2], symmetric] | 
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 290 | lemmas mcl = mult_cancel_left [THEN iffD1, THEN make_pos_rule] | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 291 | |
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 292 | (* already have this for naturals, div_mult_self1/2, but not for ints *) | 
| 65363 | 293 | lemma zdiv_mult_self: "m \<noteq> 0 \<Longrightarrow> (a + m * n) div m = a div m + n" | 
| 294 | for a m n :: int | |
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 295 | apply (rule mcl) | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 296 | prefer 2 | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 297 | apply (erule asm_rl) | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 298 | apply (simp add: zmde ring_distribs) | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 299 | done | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 300 | |
| 65363 | 301 | lemma mod_power_lem: "a > 1 \<Longrightarrow> a ^ n mod a ^ m = (if m \<le> n then 0 else a ^ n)" | 
| 302 | for a :: int | |
| 68157 | 303 | by (simp add: mod_eq_0_iff_dvd le_imp_power_dvd) | 
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 304 | |
| 65363 | 305 | lemma pl_pl_rels: "a + b = c + d \<Longrightarrow> a \<ge> c \<and> b \<le> d \<or> a \<le> c \<and> b \<ge> d" | 
| 306 | for a b c d :: nat | |
| 307 | by arith | |
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 308 | |
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
55816diff
changeset | 309 | lemmas pl_pl_rels' = add.commute [THEN [2] trans, THEN pl_pl_rels] | 
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 310 | |
| 65363 | 311 | lemma minus_eq: "m - k = m \<longleftrightarrow> k = 0 \<or> m = 0" | 
| 312 | for k m :: nat | |
| 313 | by arith | |
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 314 | |
| 65363 | 315 | lemma pl_pl_mm: "a + b = c + d \<Longrightarrow> a - c = d - b" | 
| 316 | for a b c d :: nat | |
| 317 | by arith | |
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 318 | |
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
55816diff
changeset | 319 | lemmas pl_pl_mm' = add.commute [THEN [2] trans, THEN pl_pl_mm] | 
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 320 | |
| 64245 
3d00821444fc
avoid references to lemmas designed for prover tools
 haftmann parents: 
62390diff
changeset | 321 | lemmas dme = div_mult_mod_eq | 
| 66808 
1907167b6038
elementary definition of division on natural numbers
 haftmann parents: 
66801diff
changeset | 322 | lemmas dtle = div_times_less_eq_dividend | 
| 
1907167b6038
elementary definition of division on natural numbers
 haftmann parents: 
66801diff
changeset | 323 | lemmas th2 = order_trans [OF order_refl [THEN [2] mult_le_mono] div_times_less_eq_dividend] | 
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 324 | |
| 65363 | 325 | lemma td_gal: "0 < c \<Longrightarrow> a \<ge> b * c \<longleftrightarrow> a div c \<ge> b" | 
| 326 | for a b c :: nat | |
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 327 | apply safe | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 328 | apply (erule (1) xtr4 [OF div_le_mono div_mult_self_is_m]) | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 329 | apply (erule th2) | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 330 | done | 
| 65363 | 331 | |
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 332 | lemmas td_gal_lt = td_gal [simplified not_less [symmetric], simplified] | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 333 | |
| 67120 | 334 | lemmas div_mult_le = div_times_less_eq_dividend | 
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 335 | |
| 66808 
1907167b6038
elementary definition of division on natural numbers
 haftmann parents: 
66801diff
changeset | 336 | lemmas sdl = div_nat_eqI | 
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 337 | |
| 65363 | 338 | lemma given_quot: "f > 0 \<Longrightarrow> (f * l + (f - 1)) div f = l" | 
| 339 | for f l :: nat | |
| 66808 
1907167b6038
elementary definition of division on natural numbers
 haftmann parents: 
66801diff
changeset | 340 | by (rule div_nat_eqI) (simp_all) | 
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 341 | |
| 65363 | 342 | lemma given_quot_alt: "f > 0 \<Longrightarrow> (l * f + f - Suc 0) div f = l" | 
| 343 | for f l :: nat | |
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 344 | apply (frule given_quot) | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 345 | apply (rule trans) | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 346 | prefer 2 | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 347 | apply (erule asm_rl) | 
| 65363 | 348 | apply (rule_tac f="\<lambda>n. n div f" in arg_cong) | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 349 | apply (simp add : ac_simps) | 
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 350 | done | 
| 65363 | 351 | |
| 352 | lemma diff_mod_le: "a < d \<Longrightarrow> b dvd d \<Longrightarrow> a - a mod b \<le> d - b" | |
| 353 | for a b d :: nat | |
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 354 | apply (unfold dvd_def) | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 355 | apply clarify | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 356 | apply (case_tac k) | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 357 | apply clarsimp | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 358 | apply clarify | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 359 | apply (cases "b > 0") | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
55816diff
changeset | 360 | apply (drule mult.commute [THEN xtr1]) | 
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 361 | apply (frule (1) td_gal_lt [THEN iffD1]) | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 362 | apply (clarsimp simp: le_simps) | 
| 64246 | 363 | apply (rule minus_mod_eq_mult_div [symmetric, THEN [2] xtr4]) | 
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 364 | apply (rule mult_mono) | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 365 | apply auto | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 366 | done | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 367 | |
| 65363 | 368 | lemma less_le_mult': "w * c < b * c \<Longrightarrow> 0 \<le> c \<Longrightarrow> (w + 1) * c \<le> b * c" | 
| 369 | for b c w :: int | |
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 370 | apply (rule mult_right_mono) | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 371 | apply (rule zless_imp_add1_zle) | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 372 | apply (erule (1) mult_right_less_imp_less) | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 373 | apply assumption | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 374 | done | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 375 | |
| 65363 | 376 | lemma less_le_mult: "w * c < b * c \<Longrightarrow> 0 \<le> c \<Longrightarrow> w * c + c \<le> b * c" | 
| 377 | for b c w :: int | |
| 55816 
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
 haftmann parents: 
55417diff
changeset | 378 | using less_le_mult' [of w c b] by (simp add: algebra_simps) | 
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 379 | |
| 65363 | 380 | lemmas less_le_mult_minus = iffD2 [OF le_diff_eq less_le_mult, | 
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 381 | simplified left_diff_distrib] | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 382 | |
| 65363 | 383 | lemma gen_minus: "0 < n \<Longrightarrow> f n = f (Suc (n - 1))" | 
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 384 | by auto | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 385 | |
| 65363 | 386 | lemma mpl_lem: "j \<le> i \<Longrightarrow> k < j \<Longrightarrow> i - j + k < i" | 
| 387 | for i j k :: nat | |
| 388 | by arith | |
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 389 | |
| 65363 | 390 | lemma nonneg_mod_div: "0 \<le> a \<Longrightarrow> 0 \<le> b \<Longrightarrow> 0 \<le> (a mod b) \<and> 0 \<le> a div b" | 
| 391 | for a b :: int | |
| 392 | by (cases "b = 0") (auto intro: pos_imp_zdiv_nonneg_iff [THEN iffD2]) | |
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 393 | |
| 54872 | 394 | declare iszero_0 [intro] | 
| 395 | ||
| 65363 | 396 | lemma min_pm [simp]: "min a b + (a - b) = a" | 
| 397 | for a b :: nat | |
| 54872 | 398 | by arith | 
| 65363 | 399 | |
| 400 | lemma min_pm1 [simp]: "a - b + min a b = a" | |
| 401 | for a b :: nat | |
| 54872 | 402 | by arith | 
| 403 | ||
| 65363 | 404 | lemma rev_min_pm [simp]: "min b a + (a - b) = a" | 
| 405 | for a b :: nat | |
| 54872 | 406 | by arith | 
| 407 | ||
| 65363 | 408 | lemma rev_min_pm1 [simp]: "a - b + min b a = a" | 
| 409 | for a b :: nat | |
| 54872 | 410 | by arith | 
| 411 | ||
| 65363 | 412 | lemma min_minus [simp]: "min m (m - k) = m - k" | 
| 413 | for m k :: nat | |
| 54872 | 414 | by arith | 
| 65363 | 415 | |
| 416 | lemma min_minus' [simp]: "min (m - k) m = m - k" | |
| 417 | for m k :: nat | |
| 54872 | 418 | by arith | 
| 419 | ||
| 70170 | 420 | lemmas m2pths = pos_mod_sign pos_mod_bound [OF zless2p] | 
| 421 | ||
| 422 | lemma not_B1_is_B0: "y \<noteq> 1 \<Longrightarrow> y = 0" | |
| 423 | for y :: bit | |
| 424 | by (cases y) auto | |
| 425 | ||
| 426 | lemma B1_ass_B0: | |
| 427 | fixes y :: bit | |
| 428 | assumes y: "y = 0 \<Longrightarrow> y = 1" | |
| 429 | shows "y = 1" | |
| 430 | apply (rule classical) | |
| 431 | apply (drule not_B1_is_B0) | |
| 432 | apply (erule y) | |
| 433 | done | |
| 434 | ||
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: diff
changeset | 435 | end |