| author | blanchet | 
| Thu, 28 Jul 2011 11:43:45 +0200 | |
| changeset 43996 | 4d1270ddf042 | 
| parent 38273 | d31a34569542 | 
| child 46560 | 8e252a608765 | 
| permissions | -rw-r--r-- | 
| 23664 | 1 | theory ComputeNumeral | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29668diff
changeset | 2 | imports ComputeHOL ComputeFloat | 
| 23664 | 3 | begin | 
| 4 | ||
| 5 | (* normalization of bit strings *) | |
| 26075 
815f3ccc0b45
added lemma lists {normalize,succ,pred,minus,add,mult}_bin_simps
 huffman parents: 
25919diff
changeset | 6 | lemmas bitnorm = normalize_bin_simps | 
| 23664 | 7 | |
| 8 | (* neg for bit strings *) | |
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
23664diff
changeset | 9 | lemma neg1: "neg Int.Pls = False" by (simp add: Int.Pls_def) | 
| 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
23664diff
changeset | 10 | lemma neg2: "neg Int.Min = True" apply (subst Int.Min_def) by auto | 
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
26075diff
changeset | 11 | lemma neg3: "neg (Int.Bit0 x) = neg x" apply (simp add: neg_def) apply (subst Bit0_def) by auto | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
26075diff
changeset | 12 | lemma neg4: "neg (Int.Bit1 x) = neg x" apply (simp add: neg_def) apply (subst Bit1_def) by auto | 
| 23664 | 13 | lemmas bitneg = neg1 neg2 neg3 neg4 | 
| 14 | ||
| 15 | (* iszero for bit strings *) | |
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
23664diff
changeset | 16 | lemma iszero1: "iszero Int.Pls = True" by (simp add: Int.Pls_def iszero_def) | 
| 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
23664diff
changeset | 17 | lemma iszero2: "iszero Int.Min = False" apply (subst Int.Min_def) apply (subst iszero_def) by simp | 
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
26075diff
changeset | 18 | lemma iszero3: "iszero (Int.Bit0 x) = iszero x" apply (subst Int.Bit0_def) apply (subst iszero_def)+ by auto | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
26075diff
changeset | 19 | lemma iszero4: "iszero (Int.Bit1 x) = False" apply (subst Int.Bit1_def) apply (subst iszero_def)+ apply simp by arith | 
| 23664 | 20 | lemmas bitiszero = iszero1 iszero2 iszero3 iszero4 | 
| 21 | ||
| 22 | (* lezero for bit strings *) | |
| 38273 | 23 | definition "lezero x \<longleftrightarrow> x \<le> 0" | 
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
23664diff
changeset | 24 | lemma lezero1: "lezero Int.Pls = True" unfolding Int.Pls_def lezero_def by auto | 
| 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
23664diff
changeset | 25 | lemma lezero2: "lezero Int.Min = True" unfolding Int.Min_def lezero_def by auto | 
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
26075diff
changeset | 26 | lemma lezero3: "lezero (Int.Bit0 x) = lezero x" unfolding Int.Bit0_def lezero_def by auto | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
26075diff
changeset | 27 | lemma lezero4: "lezero (Int.Bit1 x) = neg x" unfolding Int.Bit1_def lezero_def neg_def by auto | 
| 23664 | 28 | lemmas bitlezero = lezero1 lezero2 lezero3 lezero4 | 
| 29 | ||
| 30 | (* equality for bit strings *) | |
| 29037 | 31 | lemmas biteq = eq_bin_simps | 
| 23664 | 32 | |
| 33 | (* x < y for bit strings *) | |
| 29037 | 34 | lemmas bitless = less_bin_simps | 
| 23664 | 35 | |
| 36 | (* x \<le> y for bit strings *) | |
| 29037 | 37 | lemmas bitle = le_bin_simps | 
| 23664 | 38 | |
| 39 | (* succ for bit strings *) | |
| 26075 
815f3ccc0b45
added lemma lists {normalize,succ,pred,minus,add,mult}_bin_simps
 huffman parents: 
25919diff
changeset | 40 | lemmas bitsucc = succ_bin_simps | 
| 23664 | 41 | |
| 42 | (* pred for bit strings *) | |
| 26075 
815f3ccc0b45
added lemma lists {normalize,succ,pred,minus,add,mult}_bin_simps
 huffman parents: 
25919diff
changeset | 43 | lemmas bitpred = pred_bin_simps | 
| 23664 | 44 | |
| 45 | (* unary minus for bit strings *) | |
| 26075 
815f3ccc0b45
added lemma lists {normalize,succ,pred,minus,add,mult}_bin_simps
 huffman parents: 
25919diff
changeset | 46 | lemmas bituminus = minus_bin_simps | 
| 23664 | 47 | |
| 48 | (* addition for bit strings *) | |
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
26075diff
changeset | 49 | lemmas bitadd = add_bin_simps | 
| 23664 | 50 | |
| 51 | (* multiplication for bit strings *) | |
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
23664diff
changeset | 52 | lemma mult_Pls_right: "x * Int.Pls = Int.Pls" by (simp add: Pls_def) | 
| 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
23664diff
changeset | 53 | lemma mult_Min_right: "x * Int.Min = - x" by (subst mult_commute, simp add: mult_Min) | 
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
26075diff
changeset | 54 | lemma multb0x: "(Int.Bit0 x) * y = Int.Bit0 (x * y)" by (rule mult_Bit0) | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
26075diff
changeset | 55 | lemma multxb0: "x * (Int.Bit0 y) = Int.Bit0 (x * y)" unfolding Bit0_def by simp | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
26075diff
changeset | 56 | lemma multb1: "(Int.Bit1 x) * (Int.Bit1 y) = Int.Bit1 (Int.Bit0 (x * y) + x + y)" | 
| 29667 | 57 | unfolding Bit0_def Bit1_def by (simp add: algebra_simps) | 
| 23664 | 58 | lemmas bitmul = mult_Pls mult_Min mult_Pls_right mult_Min_right multb0x multxb0 multb1 | 
| 59 | ||
| 60 | lemmas bitarith = bitnorm bitiszero bitneg bitlezero biteq bitless bitle bitsucc bitpred bituminus bitadd bitmul | |
| 61 | ||
| 38273 | 62 | definition "nat_norm_number_of (x::nat) = x" | 
| 23664 | 63 | |
| 64 | lemma nat_norm_number_of: "nat_norm_number_of (number_of w) = (if lezero w then 0 else number_of w)" | |
| 65 | apply (simp add: nat_norm_number_of_def) | |
| 66 | unfolding lezero_def iszero_def neg_def | |
| 28990 | 67 | apply (simp add: numeral_simps) | 
| 23664 | 68 | done | 
| 69 | ||
| 70 | (* Normalization of nat literals *) | |
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
23664diff
changeset | 71 | lemma natnorm0: "(0::nat) = number_of (Int.Pls)" by auto | 
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
26075diff
changeset | 72 | lemma natnorm1: "(1 :: nat) = number_of (Int.Bit1 Int.Pls)" by auto | 
| 23664 | 73 | lemmas natnorm = natnorm0 natnorm1 nat_norm_number_of | 
| 74 | ||
| 75 | (* Suc *) | |
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
23664diff
changeset | 76 | lemma natsuc: "Suc (number_of x) = (if neg x then 1 else number_of (Int.succ x))" by (auto simp add: number_of_is_id) | 
| 23664 | 77 | |
| 78 | (* Addition for nat *) | |
| 79 | lemma natadd: "number_of x + ((number_of y)::nat) = (if neg x then (number_of y) else (if neg y then number_of x else (number_of (x + y))))" | |
| 29013 | 80 | unfolding nat_number_of_def number_of_is_id neg_def | 
| 81 | by auto | |
| 23664 | 82 | |
| 83 | (* Subtraction for nat *) | |
| 84 | lemma natsub: "(number_of x) - ((number_of y)::nat) = | |
| 85 | (if neg x then 0 else (if neg y then number_of x else (nat_norm_number_of (number_of (x + (- y))))))" | |
| 86 | unfolding nat_norm_number_of | |
| 87 | by (auto simp add: number_of_is_id neg_def lezero_def iszero_def Let_def nat_number_of_def) | |
| 88 | ||
| 89 | (* Multiplication for nat *) | |
| 90 | lemma natmul: "(number_of x) * ((number_of y)::nat) = | |
| 91 | (if neg x then 0 else (if neg y then 0 else number_of (x * y)))" | |
| 29013 | 92 | unfolding nat_number_of_def number_of_is_id neg_def | 
| 93 | by (simp add: nat_mult_distrib) | |
| 23664 | 94 | |
| 95 | lemma nateq: "(((number_of x)::nat) = (number_of y)) = ((lezero x \<and> lezero y) \<or> (x = y))" | |
| 96 | by (auto simp add: iszero_def lezero_def neg_def number_of_is_id) | |
| 97 | ||
| 98 | lemma natless: "(((number_of x)::nat) < (number_of y)) = ((x < y) \<and> (\<not> (lezero y)))" | |
| 29013 | 99 | by (simp add: lezero_def numeral_simps not_le) | 
| 23664 | 100 | |
| 101 | lemma natle: "(((number_of x)::nat) \<le> (number_of y)) = (y < x \<longrightarrow> lezero x)" | |
| 102 | by (auto simp add: number_of_is_id lezero_def nat_number_of_def) | |
| 103 | ||
| 104 | fun natfac :: "nat \<Rightarrow> nat" | |
| 38273 | 105 | where "natfac n = (if n = 0 then 1 else n * (natfac (n - 1)))" | 
| 23664 | 106 | |
| 107 | lemmas compute_natarith = bitarith natnorm natsuc natadd natsub natmul nateq natless natle natfac.simps | |
| 108 | ||
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
33343diff
changeset | 109 | lemma number_eq: "(((number_of x)::'a::{number_ring, linordered_idom}) = (number_of y)) = (x = y)"
 | 
| 23664 | 110 | unfolding number_of_eq | 
| 111 | apply simp | |
| 112 | done | |
| 113 | ||
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
33343diff
changeset | 114 | lemma number_le: "(((number_of x)::'a::{number_ring, linordered_idom}) \<le>  (number_of y)) = (x \<le> y)"
 | 
| 23664 | 115 | unfolding number_of_eq | 
| 116 | apply simp | |
| 117 | done | |
| 118 | ||
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
33343diff
changeset | 119 | lemma number_less: "(((number_of x)::'a::{number_ring, linordered_idom}) <  (number_of y)) = (x < y)"
 | 
| 23664 | 120 | unfolding number_of_eq | 
| 121 | apply simp | |
| 122 | done | |
| 123 | ||
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
33343diff
changeset | 124 | lemma number_diff: "((number_of x)::'a::{number_ring, linordered_idom}) - number_of y = number_of (x + (- y))"
 | 
| 23664 | 125 | apply (subst diff_number_of_eq) | 
| 126 | apply simp | |
| 127 | done | |
| 128 | ||
| 129 | lemmas number_norm = number_of_Pls[symmetric] numeral_1_eq_1[symmetric] | |
| 130 | ||
| 131 | lemmas compute_numberarith = number_of_minus[symmetric] number_of_add[symmetric] number_diff number_of_mult[symmetric] number_norm number_eq number_le number_less | |
| 132 | ||
| 133 | lemma compute_real_of_nat_number_of: "real ((number_of v)::nat) = (if neg v then 0 else number_of v)" | |
| 134 | by (simp only: real_of_nat_number_of number_of_is_id) | |
| 135 | ||
| 136 | lemma compute_nat_of_int_number_of: "nat ((number_of v)::int) = (number_of v)" | |
| 137 | by simp | |
| 138 | ||
| 139 | lemmas compute_num_conversions = compute_real_of_nat_number_of compute_nat_of_int_number_of real_number_of | |
| 140 | ||
| 141 | lemmas zpowerarith = zpower_number_of_even | |
| 142 | zpower_number_of_odd[simplified zero_eq_Numeral0_nring one_eq_Numeral1_nring] | |
| 143 | zpower_Pls zpower_Min | |
| 144 | ||
| 145 | (* div, mod *) | |
| 146 | ||
| 147 | lemma adjust: "adjust b (q, r) = (if 0 \<le> r - b then (2 * q + 1, r - b) else (2 * q, r))" | |
| 148 | by (auto simp only: adjust_def) | |
| 149 | ||
| 150 | lemma negateSnd: "negateSnd (q, r) = (q, -r)" | |
| 29657 
881f328dfbb3
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29037diff
changeset | 151 | by (simp add: negateSnd_def) | 
| 23664 | 152 | |
| 33343 | 153 | lemma divmod: "divmod_int a b = (if 0\<le>a then | 
| 23664 | 154 | if 0\<le>b then posDivAlg a b | 
| 155 | else if a=0 then (0, 0) | |
| 156 | else negateSnd (negDivAlg (-a) (-b)) | |
| 157 | else | |
| 158 | if 0<b then negDivAlg a b | |
| 159 | else negateSnd (posDivAlg (-a) (-b)))" | |
| 33343 | 160 | by (auto simp only: divmod_int_def) | 
| 23664 | 161 | |
| 33343 | 162 | lemmas compute_div_mod = div_int_def mod_int_def divmod adjust negateSnd posDivAlg.simps negDivAlg.simps | 
| 23664 | 163 | |
| 164 | ||
| 165 | ||
| 166 | (* collecting all the theorems *) | |
| 167 | ||
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
23664diff
changeset | 168 | lemma even_Pls: "even (Int.Pls) = True" | 
| 23664 | 169 | apply (unfold Pls_def even_def) | 
| 170 | by simp | |
| 171 | ||
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
23664diff
changeset | 172 | lemma even_Min: "even (Int.Min) = False" | 
| 23664 | 173 | apply (unfold Min_def even_def) | 
| 174 | by simp | |
| 175 | ||
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
26075diff
changeset | 176 | lemma even_B0: "even (Int.Bit0 x) = True" | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
26075diff
changeset | 177 | apply (unfold Bit0_def) | 
| 23664 | 178 | by simp | 
| 179 | ||
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
26075diff
changeset | 180 | lemma even_B1: "even (Int.Bit1 x) = False" | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
26075diff
changeset | 181 | apply (unfold Bit1_def) | 
| 23664 | 182 | by simp | 
| 183 | ||
| 184 | lemma even_number_of: "even ((number_of w)::int) = even w" | |
| 185 | by (simp only: number_of_is_id) | |
| 186 | ||
| 187 | lemmas compute_even = even_Pls even_Min even_B0 even_B1 even_number_of | |
| 188 | ||
| 189 | lemmas compute_numeral = compute_if compute_let compute_pair compute_bool | |
| 190 | compute_natarith compute_numberarith max_def min_def compute_num_conversions zpowerarith compute_div_mod compute_even | |
| 191 | ||
| 192 | end |