| author | wenzelm | 
| Fri, 21 May 2004 21:14:52 +0200 | |
| changeset 14766 | c0401da7726d | 
| parent 14754 | a080eeeaec14 | 
| permissions | -rw-r--r-- | 
| 1632 | 1 | (* Title: HOL/Integ/Bin.thy | 
| 6910 | 2 | ID: $Id$ | 
| 1632 | 3 | Authors: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 4 | Copyright 1994 University of Cambridge | |
| 14705 | 5 | |
| 6 | Ported from ZF to HOL by David Spelt. | |
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 7 | *) | 
| 1632 | 8 | |
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 9 | header{*Arithmetic on Binary Integers*}
 | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 10 | |
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14365diff
changeset | 11 | theory Bin = IntDef + Numeral: | 
| 1632 | 12 | |
| 14738 | 13 | axclass number_ring \<subseteq> number, comm_ring_1 | 
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 14 | number_of_Pls: "number_of bin.Pls = 0" | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 15 | number_of_Min: "number_of bin.Min = - 1" | 
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 16 | number_of_BIT: "number_of(w BIT x) = (if x then 1 else 0) + | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 17 | (number_of w) + (number_of w)" | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 18 | subsection{*Converting Numerals to Rings: @{term number_of}*}
 | 
| 1632 | 19 | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 20 | lemmas number_of = number_of_Pls number_of_Min number_of_BIT | 
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 21 | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 22 | lemma number_of_NCons [simp]: | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 23 | "number_of(NCons w b) = (number_of(w BIT b)::'a::number_ring)" | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 24 | by (induct_tac "w", simp_all add: number_of) | 
| 5491 | 25 | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 26 | lemma number_of_succ: "number_of(bin_succ w) = (1 + number_of w ::'a::number_ring)" | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 27 | apply (induct_tac "w") | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 28 | apply (simp_all add: number_of add_ac) | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 29 | done | 
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 30 | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 31 | lemma number_of_pred: "number_of(bin_pred w) = (- 1 + number_of w ::'a::number_ring)" | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 32 | apply (induct_tac "w") | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 33 | apply (simp_all add: number_of add_assoc [symmetric]) | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 34 | done | 
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 35 | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 36 | lemma number_of_minus: "number_of(bin_minus w) = (- (number_of w)::'a::number_ring)" | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 37 | apply (induct_tac "w") | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 38 | apply (simp_all del: bin_pred_Pls bin_pred_Min bin_pred_BIT | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 39 | add: number_of number_of_succ number_of_pred add_assoc) | 
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 40 | done | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 41 | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 42 | text{*This proof is complicated by the mutual recursion*}
 | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 43 | lemma number_of_add [rule_format]: | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 44 | "\<forall>w. number_of(bin_add v w) = (number_of v + number_of w::'a::number_ring)" | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 45 | apply (induct_tac "v") | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 46 | apply (simp add: number_of) | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 47 | apply (simp add: number_of number_of_pred) | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 48 | apply (rule allI) | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 49 | apply (induct_tac "w") | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 50 | apply (simp_all add: number_of bin_add_BIT_BIT number_of_succ number_of_pred add_ac) | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 51 | done | 
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 52 | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 53 | lemma number_of_mult: | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 54 | "number_of(bin_mult v w) = (number_of v * number_of w::'a::number_ring)" | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 55 | apply (induct_tac "v", simp add: number_of) | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 56 | apply (simp add: number_of number_of_minus) | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 57 | apply (simp add: number_of number_of_add left_distrib add_ac) | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 58 | done | 
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 59 | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 60 | text{*The correctness of shifting.  But it doesn't seem to give a measurable
 | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 61 | speed-up.*} | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 62 | lemma double_number_of_BIT: | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 63 | "(1+1) * number_of w = (number_of (w BIT False) ::'a::number_ring)" | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 64 | apply (induct_tac "w") | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 65 | apply (simp_all add: number_of number_of_add left_distrib add_ac) | 
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 66 | done | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 67 | |
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 68 | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 69 | text{*Converting numerals 0 and 1 to their abstract versions*}
 | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 70 | lemma numeral_0_eq_0 [simp]: "Numeral0 = (0::'a::number_ring)" | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 71 | by (simp add: number_of) | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 72 | |
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 73 | lemma numeral_1_eq_1 [simp]: "Numeral1 = (1::'a::number_ring)" | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 74 | by (simp add: number_of) | 
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 75 | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 76 | text{*Special-case simplification for small constants*}
 | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 77 | |
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 78 | text{*Unary minus for the abstract constant 1. Cannot be inserted
 | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 79 |   as a simprule until later: it is @{text number_of_Min} re-oriented!*}
 | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 80 | lemma numeral_m1_eq_minus_1: "(-1::'a::number_ring) = - 1" | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 81 | by (simp add: number_of) | 
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 82 | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 83 | lemma mult_minus1 [simp]: "-1 * z = -(z::'a::number_ring)" | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 84 | by (simp add: numeral_m1_eq_minus_1) | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 85 | |
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 86 | lemma mult_minus1_right [simp]: "z * -1 = -(z::'a::number_ring)" | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 87 | by (simp add: numeral_m1_eq_minus_1) | 
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 88 | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 89 | (*Negation of a coefficient*) | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 90 | lemma minus_number_of_mult [simp]: | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 91 | "- (number_of w) * z = number_of(bin_minus w) * (z::'a::number_ring)" | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 92 | by (simp add: number_of_minus) | 
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 93 | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 94 | text{*Subtraction*}
 | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 95 | lemma diff_number_of_eq: | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 96 | "number_of v - number_of w = | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 97 | (number_of(bin_add v (bin_minus w))::'a::number_ring)" | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 98 | by (simp add: diff_minus number_of_add number_of_minus) | 
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 99 | |
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 100 | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 101 | subsection{*Equality of Binary Numbers*}
 | 
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 102 | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 103 | text{*First version by Norbert Voelker*}
 | 
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 104 | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 105 | lemma eq_number_of_eq: | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 106 | "((number_of x::'a::number_ring) = number_of y) = | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 107 | iszero (number_of (bin_add x (bin_minus y)) :: 'a)" | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 108 | by (simp add: iszero_def compare_rls number_of_add number_of_minus) | 
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 109 | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 110 | lemma iszero_number_of_Pls: "iszero ((number_of bin.Pls)::'a::number_ring)" | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 111 | by (simp add: iszero_def numeral_0_eq_0) | 
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 112 | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 113 | lemma nonzero_number_of_Min: "~ iszero ((number_of bin.Min)::'a::number_ring)" | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 114 | by (simp add: iszero_def numeral_m1_eq_minus_1 eq_commute) | 
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 115 | |
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 116 | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 117 | subsection{*Comparisons, for Ordered Rings*}
 | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 118 | |
| 14738 | 119 | lemma double_eq_0_iff: "(a + a = 0) = (a = (0::'a::ordered_idom))" | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 120 | proof - | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 121 | have "a + a = (1+1)*a" by (simp add: left_distrib) | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 122 | with zero_less_two [where 'a = 'a] | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 123 | show ?thesis by force | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 124 | qed | 
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 125 | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 126 | lemma le_imp_0_less: | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 127 | assumes le: "0 \<le> z" shows "(0::int) < 1 + z" | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 128 | proof - | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 129 | have "0 \<le> z" . | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 130 | also have "... < z + 1" by (rule less_add_one) | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 131 | also have "... = 1 + z" by (simp add: add_ac) | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 132 | finally show "0 < 1 + z" . | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 133 | qed | 
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 134 | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 135 | lemma odd_nonzero: "1 + z + z \<noteq> (0::int)"; | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 136 | proof (cases z rule: int_cases) | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 137 | case (nonneg n) | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14473diff
changeset | 138 | have le: "0 \<le> z+z" by (simp add: nonneg add_increasing) | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 139 | thus ?thesis using le_imp_0_less [OF le] | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 140 | by (auto simp add: add_assoc) | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 141 | next | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 142 | case (neg n) | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 143 | show ?thesis | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 144 | proof | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 145 | assume eq: "1 + z + z = 0" | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 146 | have "0 < 1 + (int n + int n)" | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 147 | by (simp add: le_imp_0_less add_increasing) | 
| 14754 
a080eeeaec14
Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
 obua parents: 
14738diff
changeset | 148 | also have "... = - (1 + z + z)" | 
| 
a080eeeaec14
Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
 obua parents: 
14738diff
changeset | 149 | by (simp add: neg add_assoc [symmetric]) | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 150 | also have "... = 0" by (simp add: eq) | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 151 | finally have "0<0" .. | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 152 | thus False by blast | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 153 | qed | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 154 | qed | 
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 155 | |
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 156 | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 157 | text{*The premise involving @{term Ints} prevents @{term "a = 1/2"}.*}
 | 
| 14738 | 158 | lemma Ints_odd_nonzero: "a \<in> Ints ==> 1 + a + a \<noteq> (0::'a::ordered_idom)" | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 159 | proof (unfold Ints_def) | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 160 | assume "a \<in> range of_int" | 
| 14473 | 161 | then obtain z where a: "a = of_int z" .. | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 162 | show ?thesis | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 163 | proof | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 164 | assume eq: "1 + a + a = 0" | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 165 | hence "of_int (1 + z + z) = (of_int 0 :: 'a)" by (simp add: a) | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 166 | hence "1 + z + z = 0" by (simp only: of_int_eq_iff) | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 167 | with odd_nonzero show False by blast | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 168 | qed | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 169 | qed | 
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 170 | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 171 | lemma Ints_number_of: "(number_of w :: 'a::number_ring) \<in> Ints" | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 172 | by (induct_tac "w", simp_all add: number_of) | 
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 173 | |
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 174 | lemma iszero_number_of_BIT: | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 175 | "iszero (number_of (w BIT x)::'a) = | 
| 14738 | 176 |       (~x & iszero (number_of w::'a::{ordered_idom,number_ring}))"
 | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 177 | by (simp add: iszero_def compare_rls zero_reorient double_eq_0_iff | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 178 | number_of Ints_odd_nonzero [OF Ints_number_of]) | 
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 179 | |
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14365diff
changeset | 180 | lemma iszero_number_of_0: | 
| 14738 | 181 |      "iszero (number_of (w BIT False) :: 'a::{ordered_idom,number_ring}) = 
 | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 182 | iszero (number_of w :: 'a)" | 
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 183 | by (simp only: iszero_number_of_BIT simp_thms) | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 184 | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 185 | lemma iszero_number_of_1: | 
| 14738 | 186 |      "~ iszero (number_of (w BIT True)::'a::{ordered_idom,number_ring})"
 | 
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 187 | by (simp only: iszero_number_of_BIT simp_thms) | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 188 | |
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 189 | |
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 190 | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 191 | subsection{*The Less-Than Relation*}
 | 
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 192 | |
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 193 | lemma less_number_of_eq_neg: | 
| 14738 | 194 |     "((number_of x::'a::{ordered_idom,number_ring}) < number_of y)
 | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 195 | = neg (number_of (bin_add x (bin_minus y)) :: 'a)" | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 196 | apply (subst less_iff_diff_less_0) | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 197 | apply (simp add: neg_def diff_minus number_of_add number_of_minus) | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 198 | done | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 199 | |
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 200 | text{*If @{term Numeral0} is rewritten to 0 then this rule can't be applied:
 | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 201 |   @{term Numeral0} IS @{term "number_of Pls"} *}
 | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 202 | lemma not_neg_number_of_Pls: | 
| 14738 | 203 |      "~ neg (number_of bin.Pls ::'a::{ordered_idom,number_ring})"
 | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 204 | by (simp add: neg_def numeral_0_eq_0) | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 205 | |
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 206 | lemma neg_number_of_Min: | 
| 14738 | 207 |      "neg (number_of bin.Min ::'a::{ordered_idom,number_ring})"
 | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 208 | by (simp add: neg_def zero_less_one numeral_m1_eq_minus_1) | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 209 | |
| 14738 | 210 | lemma double_less_0_iff: "(a + a < 0) = (a < (0::'a::ordered_idom))" | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 211 | proof - | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 212 | have "(a + a < 0) = ((1+1)*a < 0)" by (simp add: left_distrib) | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 213 | also have "... = (a < 0)" | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 214 | by (simp add: mult_less_0_iff zero_less_two | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 215 | order_less_not_sym [OF zero_less_two]) | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 216 | finally show ?thesis . | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 217 | qed | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 218 | |
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 219 | lemma odd_less_0: "(1 + z + z < 0) = (z < (0::int))"; | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 220 | proof (cases z rule: int_cases) | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 221 | case (nonneg n) | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 222 | thus ?thesis by (simp add: linorder_not_less add_assoc add_increasing | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 223 | le_imp_0_less [THEN order_less_imp_le]) | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 224 | next | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 225 | case (neg n) | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 226 | thus ?thesis by (simp del: int_Suc | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 227 | add: int_Suc0_eq_1 [symmetric] zadd_int compare_rls) | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 228 | qed | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 229 | |
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 230 | text{*The premise involving @{term Ints} prevents @{term "a = 1/2"}.*}
 | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 231 | lemma Ints_odd_less_0: | 
| 14738 | 232 | "a \<in> Ints ==> (1 + a + a < 0) = (a < (0::'a::ordered_idom))"; | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 233 | proof (unfold Ints_def) | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 234 | assume "a \<in> range of_int" | 
| 14473 | 235 | then obtain z where a: "a = of_int z" .. | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 236 | hence "((1::'a) + a + a < 0) = (of_int (1 + z + z) < (of_int 0 :: 'a))" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14473diff
changeset | 237 | by (simp add: a) | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 238 | also have "... = (z < 0)" by (simp only: of_int_less_iff odd_less_0) | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14473diff
changeset | 239 | also have "... = (a < 0)" by (simp add: a) | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 240 | finally show ?thesis . | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 241 | qed | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 242 | |
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 243 | lemma neg_number_of_BIT: | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 244 | "neg (number_of (w BIT x)::'a) = | 
| 14738 | 245 |       neg (number_of w :: 'a::{ordered_idom,number_ring})"
 | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 246 | by (simp add: number_of neg_def double_less_0_iff | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 247 | Ints_odd_less_0 [OF Ints_number_of]) | 
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 248 | |
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 249 | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 250 | text{*Less-Than or Equals*}
 | 
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 251 | |
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14288diff
changeset | 252 | text{*Reduces @{term "a\<le>b"} to @{term "~ (b<a)"} for ALL numerals*}
 | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14288diff
changeset | 253 | lemmas le_number_of_eq_not_less = | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 254 | linorder_not_less [of "number_of w" "number_of v", symmetric, | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 255 | standard] | 
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14288diff
changeset | 256 | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 257 | lemma le_number_of_eq: | 
| 14738 | 258 |     "((number_of x::'a::{ordered_idom,number_ring}) \<le> number_of y)
 | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 259 | = (~ (neg (number_of (bin_add y (bin_minus x)) :: 'a)))" | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 260 | by (simp add: le_number_of_eq_not_less less_number_of_eq_neg) | 
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14288diff
changeset | 261 | |
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 262 | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 263 | text{*Absolute value (@{term abs})*}
 | 
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 264 | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 265 | lemma abs_number_of: | 
| 14738 | 266 |      "abs(number_of x::'a::{ordered_idom,number_ring}) =
 | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 267 | (if number_of x < (0::'a) then -number_of x else number_of x)" | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 268 | by (simp add: abs_if) | 
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 269 | |
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 270 | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 271 | text{*Re-orientation of the equation nnn=x*}
 | 
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 272 | lemma number_of_reorient: "(number_of w = x) = (x = number_of w)" | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 273 | by auto | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 274 | |
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 275 | |
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 276 | (*Delete the original rewrites, with their clumsy conditional expressions*) | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 277 | declare bin_succ_BIT [simp del] bin_pred_BIT [simp del] | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 278 | bin_minus_BIT [simp del] | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 279 | |
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 280 | declare bin_add_BIT [simp del] bin_mult_BIT [simp del] | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 281 | declare NCons_Pls [simp del] NCons_Min [simp del] | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 282 | |
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 283 | (*Hide the binary representation of integer constants*) | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 284 | declare number_of_Pls [simp del] number_of_Min [simp del] | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 285 | number_of_BIT [simp del] | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 286 | |
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 287 | |
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 288 | |
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 289 | (*Simplification of arithmetic operations on integer constants. | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 290 | Note that bin_pred_Pls, etc. were added to the simpset by primrec.*) | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 291 | |
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 292 | lemmas NCons_simps = NCons_Pls_0 NCons_Pls_1 NCons_Min_0 NCons_Min_1 NCons_BIT | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 293 | |
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 294 | lemmas bin_arith_extra_simps = | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 295 | number_of_add [symmetric] | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 296 | number_of_minus [symmetric] numeral_m1_eq_minus_1 [symmetric] | 
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 297 | number_of_mult [symmetric] | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 298 | bin_succ_1 bin_succ_0 | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 299 | bin_pred_1 bin_pred_0 | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 300 | bin_minus_1 bin_minus_0 | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 301 | bin_add_Pls_right bin_add_Min_right | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 302 | bin_add_BIT_0 bin_add_BIT_10 bin_add_BIT_11 | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 303 | diff_number_of_eq abs_number_of abs_zero abs_one | 
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 304 | bin_mult_1 bin_mult_0 NCons_simps | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 305 | |
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 306 | (*For making a minimal simpset, one must include these default simprules | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 307 | of thy. Also include simp_thms, or at least (~False)=True*) | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 308 | lemmas bin_arith_simps = | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 309 | bin_pred_Pls bin_pred_Min | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 310 | bin_succ_Pls bin_succ_Min | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 311 | bin_add_Pls bin_add_Min | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 312 | bin_minus_Pls bin_minus_Min | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 313 | bin_mult_Pls bin_mult_Min bin_arith_extra_simps | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 314 | |
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 315 | (*Simplification of relational operations*) | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 316 | lemmas bin_rel_simps = | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 317 | eq_number_of_eq iszero_number_of_Pls nonzero_number_of_Min | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 318 | iszero_number_of_0 iszero_number_of_1 | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 319 | less_number_of_eq_neg | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 320 | not_neg_number_of_Pls not_neg_0 not_neg_1 not_iszero_1 | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 321 | neg_number_of_Min neg_number_of_BIT | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 322 | le_number_of_eq | 
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 323 | |
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 324 | declare bin_arith_extra_simps [simp] | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 325 | declare bin_rel_simps [simp] | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 326 | |
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 327 | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 328 | subsection{*Simplification of arithmetic when nested to the right*}
 | 
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 329 | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 330 | lemma add_number_of_left [simp]: | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 331 | "number_of v + (number_of w + z) = | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 332 | (number_of(bin_add v w) + z::'a::number_ring)" | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 333 | by (simp add: add_assoc [symmetric]) | 
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 334 | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 335 | lemma mult_number_of_left [simp]: | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 336 | "number_of v * (number_of w * z) = | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 337 | (number_of(bin_mult v w) * z::'a::number_ring)" | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 338 | by (simp add: mult_assoc [symmetric]) | 
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 339 | |
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 340 | lemma add_number_of_diff1: | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 341 | "number_of v + (number_of w - c) = | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 342 | number_of(bin_add v w) - (c::'a::number_ring)" | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 343 | by (simp add: diff_minus add_number_of_left) | 
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 344 | |
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 345 | lemma add_number_of_diff2 [simp]: "number_of v + (c - number_of w) = | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 346 | number_of (bin_add v (bin_minus w)) + (c::'a::number_ring)" | 
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 347 | apply (subst diff_number_of_eq [symmetric]) | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 348 | apply (simp only: compare_rls) | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 349 | done | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
13491diff
changeset | 350 | |
| 1632 | 351 | end |