haftmann@25919: (* Title: Int.thy haftmann@25919: ID: $Id$ haftmann@25919: Author: Lawrence C Paulson, Cambridge University Computer Laboratory haftmann@25919: Tobias Nipkow, Florian Haftmann, TU Muenchen haftmann@25919: Copyright 1994 University of Cambridge haftmann@25919: haftmann@25919: *) haftmann@25919: haftmann@25919: header {* The Integers as Equivalence Classes over Pairs of Natural Numbers *} haftmann@25919: haftmann@25919: theory Int krauss@26748: imports Equiv_Relations Nat Wellfounded haftmann@25919: uses haftmann@25919: ("Tools/numeral.ML") haftmann@25919: ("Tools/numeral_syntax.ML") haftmann@25919: ("~~/src/Provers/Arith/assoc_fold.ML") haftmann@25919: "~~/src/Provers/Arith/cancel_numerals.ML" haftmann@25919: "~~/src/Provers/Arith/combine_numerals.ML" haftmann@28952: ("Tools/int_arith.ML") haftmann@25919: begin haftmann@25919: haftmann@25919: subsection {* The equivalence relation underlying the integers *} haftmann@25919: haftmann@28661: definition intrel :: "((nat \ nat) \ (nat \ nat)) set" where haftmann@28562: [code del]: "intrel = {((x, y), (u, v)) | x y u v. x + v = u +y }" haftmann@25919: haftmann@25919: typedef (Integ) haftmann@25919: int = "UNIV//intrel" haftmann@25919: by (auto simp add: quotient_def) haftmann@25919: haftmann@25919: instantiation int :: "{zero, one, plus, minus, uminus, times, ord, abs, sgn}" haftmann@25919: begin haftmann@25919: haftmann@25919: definition haftmann@28562: Zero_int_def [code del]: "0 = Abs_Integ (intrel `` {(0, 0)})" haftmann@25919: haftmann@25919: definition haftmann@28562: One_int_def [code del]: "1 = Abs_Integ (intrel `` {(1, 0)})" haftmann@25919: haftmann@25919: definition haftmann@28562: add_int_def [code del]: "z + w = Abs_Integ haftmann@25919: (\(x, y) \ Rep_Integ z. \(u, v) \ Rep_Integ w. haftmann@25919: intrel `` {(x + u, y + v)})" haftmann@25919: haftmann@25919: definition haftmann@28562: minus_int_def [code del]: haftmann@25919: "- z = Abs_Integ (\(x, y) \ Rep_Integ z. intrel `` {(y, x)})" haftmann@25919: haftmann@25919: definition haftmann@28562: diff_int_def [code del]: "z - w = z + (-w \ int)" haftmann@25919: haftmann@25919: definition haftmann@28562: mult_int_def [code del]: "z * w = Abs_Integ haftmann@25919: (\(x, y) \ Rep_Integ z. \(u,v ) \ Rep_Integ w. haftmann@25919: intrel `` {(x*u + y*v, x*v + y*u)})" haftmann@25919: haftmann@25919: definition haftmann@28562: le_int_def [code del]: haftmann@25919: "z \ w \ (\x y u v. x+v \ u+y \ (x, y) \ Rep_Integ z \ (u, v) \ Rep_Integ w)" haftmann@25919: haftmann@25919: definition haftmann@28562: less_int_def [code del]: "(z\int) < w \ z \ w \ z \ w" haftmann@25919: haftmann@25919: definition haftmann@25919: zabs_def: "\i\int\ = (if i < 0 then - i else i)" haftmann@25919: haftmann@25919: definition haftmann@25919: zsgn_def: "sgn (i\int) = (if i=0 then 0 else if 0 intrel) = (x+v = u+y)" haftmann@25919: by (simp add: intrel_def) haftmann@25919: haftmann@25919: lemma equiv_intrel: "equiv UNIV intrel" haftmann@25919: by (simp add: intrel_def equiv_def refl_def sym_def trans_def) haftmann@25919: haftmann@25919: text{*Reduces equality of equivalence classes to the @{term intrel} relation: haftmann@25919: @{term "(intrel `` {x} = intrel `` {y}) = ((x,y) \ intrel)"} *} haftmann@25919: lemmas equiv_intrel_iff [simp] = eq_equiv_class_iff [OF equiv_intrel UNIV_I UNIV_I] haftmann@25919: haftmann@25919: text{*All equivalence classes belong to set of representatives*} haftmann@25919: lemma [simp]: "intrel``{(x,y)} \ Integ" haftmann@25919: by (auto simp add: Integ_def intrel_def quotient_def) haftmann@25919: haftmann@25919: text{*Reduces equality on abstractions to equality on representatives: haftmann@25919: @{prop "\x \ Integ; y \ Integ\ \ (Abs_Integ x = Abs_Integ y) = (x=y)"} *} haftmann@25919: declare Abs_Integ_inject [simp,noatp] Abs_Integ_inverse [simp,noatp] haftmann@25919: haftmann@25919: text{*Case analysis on the representation of an integer as an equivalence haftmann@25919: class of pairs of naturals.*} haftmann@25919: lemma eq_Abs_Integ [case_names Abs_Integ, cases type: int]: haftmann@25919: "(!!x y. z = Abs_Integ(intrel``{(x,y)}) ==> P) ==> P" haftmann@25919: apply (rule Abs_Integ_cases [of z]) haftmann@25919: apply (auto simp add: Integ_def quotient_def) haftmann@25919: done haftmann@25919: haftmann@25919: haftmann@25919: subsection {* Arithmetic Operations *} haftmann@25919: haftmann@25919: lemma minus: "- Abs_Integ(intrel``{(x,y)}) = Abs_Integ(intrel `` {(y,x)})" haftmann@25919: proof - haftmann@25919: have "(\(x,y). intrel``{(y,x)}) respects intrel" haftmann@25919: by (simp add: congruent_def) haftmann@25919: thus ?thesis haftmann@25919: by (simp add: minus_int_def UN_equiv_class [OF equiv_intrel]) haftmann@25919: qed haftmann@25919: haftmann@25919: lemma add: haftmann@25919: "Abs_Integ (intrel``{(x,y)}) + Abs_Integ (intrel``{(u,v)}) = haftmann@25919: Abs_Integ (intrel``{(x+u, y+v)})" haftmann@25919: proof - haftmann@25919: have "(\z w. (\(x,y). (\(u,v). intrel `` {(x+u, y+v)}) w) z) haftmann@25919: respects2 intrel" haftmann@25919: by (simp add: congruent2_def) haftmann@25919: thus ?thesis haftmann@25919: by (simp add: add_int_def UN_UN_split_split_eq haftmann@25919: UN_equiv_class2 [OF equiv_intrel equiv_intrel]) haftmann@25919: qed haftmann@25919: haftmann@25919: text{*Congruence property for multiplication*} haftmann@25919: lemma mult_congruent2: haftmann@25919: "(%p1 p2. (%(x,y). (%(u,v). intrel``{(x*u + y*v, x*v + y*u)}) p2) p1) haftmann@25919: respects2 intrel" haftmann@25919: apply (rule equiv_intrel [THEN congruent2_commuteI]) haftmann@25919: apply (force simp add: mult_ac, clarify) haftmann@25919: apply (simp add: congruent_def mult_ac) haftmann@25919: apply (rename_tac u v w x y z) haftmann@25919: apply (subgoal_tac "u*y + x*y = w*y + v*y & u*z + x*z = w*z + v*z") haftmann@25919: apply (simp add: mult_ac) haftmann@25919: apply (simp add: add_mult_distrib [symmetric]) haftmann@25919: done haftmann@25919: haftmann@25919: lemma mult: haftmann@25919: "Abs_Integ((intrel``{(x,y)})) * Abs_Integ((intrel``{(u,v)})) = haftmann@25919: Abs_Integ(intrel `` {(x*u + y*v, x*v + y*u)})" haftmann@25919: by (simp add: mult_int_def UN_UN_split_split_eq mult_congruent2 haftmann@25919: UN_equiv_class2 [OF equiv_intrel equiv_intrel]) haftmann@25919: haftmann@25919: text{*The integers form a @{text comm_ring_1}*} haftmann@25919: instance int :: comm_ring_1 haftmann@25919: proof haftmann@25919: fix i j k :: int haftmann@25919: show "(i + j) + k = i + (j + k)" haftmann@25919: by (cases i, cases j, cases k) (simp add: add add_assoc) haftmann@25919: show "i + j = j + i" haftmann@25919: by (cases i, cases j) (simp add: add_ac add) haftmann@25919: show "0 + i = i" haftmann@25919: by (cases i) (simp add: Zero_int_def add) haftmann@25919: show "- i + i = 0" haftmann@25919: by (cases i) (simp add: Zero_int_def minus add) haftmann@25919: show "i - j = i + - j" haftmann@25919: by (simp add: diff_int_def) haftmann@25919: show "(i * j) * k = i * (j * k)" haftmann@25919: by (cases i, cases j, cases k) (simp add: mult ring_simps) haftmann@25919: show "i * j = j * i" haftmann@25919: by (cases i, cases j) (simp add: mult ring_simps) haftmann@25919: show "1 * i = i" haftmann@25919: by (cases i) (simp add: One_int_def mult) haftmann@25919: show "(i + j) * k = i * k + j * k" haftmann@25919: by (cases i, cases j, cases k) (simp add: add mult ring_simps) haftmann@25919: show "0 \ (1::int)" haftmann@25919: by (simp add: Zero_int_def One_int_def) haftmann@25919: qed haftmann@25919: haftmann@25919: lemma int_def: "of_nat m = Abs_Integ (intrel `` {(m, 0)})" haftmann@25919: by (induct m, simp_all add: Zero_int_def One_int_def add) haftmann@25919: haftmann@25919: haftmann@25919: subsection {* The @{text "\"} Ordering *} haftmann@25919: haftmann@25919: lemma le: haftmann@25919: "(Abs_Integ(intrel``{(x,y)}) \ Abs_Integ(intrel``{(u,v)})) = (x+v \ u+y)" haftmann@25919: by (force simp add: le_int_def) haftmann@25919: haftmann@25919: lemma less: haftmann@25919: "(Abs_Integ(intrel``{(x,y)}) < Abs_Integ(intrel``{(u,v)})) = (x+v < u+y)" haftmann@25919: by (simp add: less_int_def le order_less_le) haftmann@25919: haftmann@25919: instance int :: linorder haftmann@25919: proof haftmann@25919: fix i j k :: int haftmann@27682: show antisym: "i \ j \ j \ i \ i = j" haftmann@27682: by (cases i, cases j) (simp add: le) haftmann@27682: show "(i < j) = (i \ j \ \ j \ i)" haftmann@27682: by (auto simp add: less_int_def dest: antisym) haftmann@25919: show "i \ i" haftmann@25919: by (cases i) (simp add: le) haftmann@25919: show "i \ j \ j \ k \ i \ k" haftmann@25919: by (cases i, cases j, cases k) (simp add: le) haftmann@25919: show "i \ j \ j \ i" haftmann@25919: by (cases i, cases j) (simp add: le linorder_linear) haftmann@25919: qed haftmann@25919: haftmann@25919: instantiation int :: distrib_lattice haftmann@25919: begin haftmann@25919: haftmann@25919: definition haftmann@25919: "(inf \ int \ int \ int) = min" haftmann@25919: haftmann@25919: definition haftmann@25919: "(sup \ int \ int \ int) = max" haftmann@25919: haftmann@25919: instance haftmann@25919: by intro_classes haftmann@25919: (auto simp add: inf_int_def sup_int_def min_max.sup_inf_distrib1) haftmann@25919: haftmann@25919: end haftmann@25919: haftmann@25919: instance int :: pordered_cancel_ab_semigroup_add haftmann@25919: proof haftmann@25919: fix i j k :: int haftmann@25919: show "i \ j \ k + i \ k + j" haftmann@25919: by (cases i, cases j, cases k) (simp add: le add) haftmann@25919: qed haftmann@25919: haftmann@25961: haftmann@25919: text{*Strict Monotonicity of Multiplication*} haftmann@25919: haftmann@25919: text{*strict, in 1st argument; proof is by induction on k>0*} haftmann@25919: lemma zmult_zless_mono2_lemma: haftmann@25919: "(i::int) 0 of_nat k * i < of_nat k * j" haftmann@25919: apply (induct "k", simp) haftmann@25919: apply (simp add: left_distrib) haftmann@25919: apply (case_tac "k=0") haftmann@25919: apply (simp_all add: add_strict_mono) haftmann@25919: done haftmann@25919: haftmann@25919: lemma zero_le_imp_eq_int: "(0::int) \ k ==> \n. k = of_nat n" haftmann@25919: apply (cases k) haftmann@25919: apply (auto simp add: le add int_def Zero_int_def) haftmann@25919: apply (rule_tac x="x-y" in exI, simp) haftmann@25919: done haftmann@25919: haftmann@25919: lemma zero_less_imp_eq_int: "(0::int) < k ==> \n>0. k = of_nat n" haftmann@25919: apply (cases k) haftmann@25919: apply (simp add: less int_def Zero_int_def) haftmann@25919: apply (rule_tac x="x-y" in exI, simp) haftmann@25919: done haftmann@25919: haftmann@25919: lemma zmult_zless_mono2: "[| i k*i < k*j" haftmann@25919: apply (drule zero_less_imp_eq_int) haftmann@25919: apply (auto simp add: zmult_zless_mono2_lemma) haftmann@25919: done haftmann@25919: haftmann@25919: text{*The integers form an ordered integral domain*} haftmann@25919: instance int :: ordered_idom haftmann@25919: proof haftmann@25919: fix i j k :: int haftmann@25919: show "i < j \ 0 < k \ k * i < k * j" haftmann@25919: by (rule zmult_zless_mono2) haftmann@25919: show "\i\ = (if i < 0 then -i else i)" haftmann@25919: by (simp only: zabs_def) haftmann@25919: show "sgn (i\int) = (if i=0 then 0 else if 0 w + (1\int) \ z" haftmann@25919: apply (cases w, cases z) haftmann@25919: apply (simp add: less le add One_int_def) haftmann@25919: done haftmann@25919: haftmann@25919: lemma zless_iff_Suc_zadd: haftmann@25919: "(w \ int) < z \ (\n. z = w + of_nat (Suc n))" haftmann@25919: apply (cases z, cases w) haftmann@25919: apply (auto simp add: less add int_def) haftmann@25919: apply (rename_tac a b c d) haftmann@25919: apply (rule_tac x="a+d - Suc(c+b)" in exI) haftmann@25919: apply arith haftmann@25919: done haftmann@25919: haftmann@25919: lemmas int_distrib = haftmann@25919: left_distrib [of "z1::int" "z2" "w", standard] haftmann@25919: right_distrib [of "w::int" "z1" "z2", standard] haftmann@25919: left_diff_distrib [of "z1::int" "z2" "w", standard] haftmann@25919: right_diff_distrib [of "w::int" "z1" "z2", standard] haftmann@25919: haftmann@25919: haftmann@25919: subsection {* Embedding of the Integers into any @{text ring_1}: @{text of_int}*} haftmann@25919: haftmann@25919: context ring_1 haftmann@25919: begin haftmann@25919: haftmann@25919: definition haftmann@25919: of_int :: "int \ 'a" haftmann@25919: where haftmann@28562: [code del]: "of_int z = contents (\(i, j) \ Rep_Integ z. { of_nat i - of_nat j })" haftmann@25919: haftmann@25919: lemma of_int: "of_int (Abs_Integ (intrel `` {(i,j)})) = of_nat i - of_nat j" haftmann@25919: proof - haftmann@25919: have "(\(i,j). { of_nat i - (of_nat j :: 'a) }) respects intrel" haftmann@25919: by (simp add: congruent_def compare_rls of_nat_add [symmetric] haftmann@25919: del: of_nat_add) haftmann@25919: thus ?thesis haftmann@25919: by (simp add: of_int_def UN_equiv_class [OF equiv_intrel]) haftmann@25919: qed haftmann@25919: haftmann@25919: lemma of_int_0 [simp]: "of_int 0 = 0" haftmann@25919: by (simp add: of_int Zero_int_def) haftmann@25919: haftmann@25919: lemma of_int_1 [simp]: "of_int 1 = 1" haftmann@25919: by (simp add: of_int One_int_def) haftmann@25919: haftmann@25919: lemma of_int_add [simp]: "of_int (w+z) = of_int w + of_int z" haftmann@25919: by (cases w, cases z, simp add: compare_rls of_int OrderedGroup.compare_rls add) haftmann@25919: haftmann@25919: lemma of_int_minus [simp]: "of_int (-z) = - (of_int z)" haftmann@25919: by (cases z, simp add: compare_rls of_int minus) haftmann@25919: haftmann@25919: lemma of_int_diff [simp]: "of_int (w - z) = of_int w - of_int z" haftmann@25919: by (simp add: OrderedGroup.diff_minus diff_minus) haftmann@25919: haftmann@25919: lemma of_int_mult [simp]: "of_int (w*z) = of_int w * of_int z" haftmann@25919: apply (cases w, cases z) haftmann@25919: apply (simp add: compare_rls of_int left_diff_distrib right_diff_distrib haftmann@25919: mult add_ac of_nat_mult) haftmann@25919: done haftmann@25919: haftmann@25919: text{*Collapse nested embeddings*} haftmann@25919: lemma of_int_of_nat_eq [simp]: "of_int (of_nat n) = of_nat n" haftmann@25919: by (induct n) auto haftmann@25919: haftmann@25919: end haftmann@25919: haftmann@25919: context ordered_idom haftmann@25919: begin haftmann@25919: haftmann@25919: lemma of_int_le_iff [simp]: haftmann@25919: "of_int w \ of_int z \ w \ z" haftmann@25919: by (cases w, cases z, simp add: of_int le minus compare_rls of_nat_add [symmetric] del: of_nat_add) haftmann@25919: haftmann@25919: text{*Special cases where either operand is zero*} haftmann@25919: lemmas of_int_0_le_iff [simp] = of_int_le_iff [of 0, simplified] haftmann@25919: lemmas of_int_le_0_iff [simp] = of_int_le_iff [of _ 0, simplified] haftmann@25919: haftmann@25919: lemma of_int_less_iff [simp]: haftmann@25919: "of_int w < of_int z \ w < z" haftmann@25919: by (simp add: not_le [symmetric] linorder_not_le [symmetric]) haftmann@25919: haftmann@25919: text{*Special cases where either operand is zero*} haftmann@25919: lemmas of_int_0_less_iff [simp] = of_int_less_iff [of 0, simplified] haftmann@25919: lemmas of_int_less_0_iff [simp] = of_int_less_iff [of _ 0, simplified] haftmann@25919: haftmann@25919: end haftmann@25919: haftmann@25919: text{*Class for unital rings with characteristic zero. haftmann@25919: Includes non-ordered rings like the complex numbers.*} haftmann@25919: class ring_char_0 = ring_1 + semiring_char_0 haftmann@25919: begin haftmann@25919: haftmann@25919: lemma of_int_eq_iff [simp]: haftmann@25919: "of_int w = of_int z \ w = z" haftmann@25919: apply (cases w, cases z, simp add: of_int) haftmann@25919: apply (simp only: diff_eq_eq diff_add_eq eq_diff_eq) haftmann@25919: apply (simp only: of_nat_add [symmetric] of_nat_eq_iff) haftmann@25919: done haftmann@25919: haftmann@25919: text{*Special cases where either operand is zero*} haftmann@25919: lemmas of_int_0_eq_iff [simp] = of_int_eq_iff [of 0, simplified] haftmann@25919: lemmas of_int_eq_0_iff [simp] = of_int_eq_iff [of _ 0, simplified] haftmann@25919: haftmann@25919: end haftmann@25919: haftmann@25919: text{*Every @{text ordered_idom} has characteristic zero.*} haftmann@25919: subclass (in ordered_idom) ring_char_0 by intro_locales haftmann@25919: haftmann@25919: lemma of_int_eq_id [simp]: "of_int = id" haftmann@25919: proof haftmann@25919: fix z show "of_int z = id z" haftmann@25919: by (cases z) (simp add: of_int add minus int_def diff_minus) haftmann@25919: qed haftmann@25919: haftmann@25919: haftmann@25919: subsection {* Magnitude of an Integer, as a Natural Number: @{text nat} *} haftmann@25919: haftmann@25919: definition haftmann@25919: nat :: "int \ nat" haftmann@25919: where haftmann@28562: [code del]: "nat z = contents (\(x, y) \ Rep_Integ z. {x-y})" haftmann@25919: haftmann@25919: lemma nat: "nat (Abs_Integ (intrel``{(x,y)})) = x-y" haftmann@25919: proof - haftmann@25919: have "(\(x,y). {x-y}) respects intrel" haftmann@25919: by (simp add: congruent_def) arith haftmann@25919: thus ?thesis haftmann@25919: by (simp add: nat_def UN_equiv_class [OF equiv_intrel]) haftmann@25919: qed haftmann@25919: haftmann@25919: lemma nat_int [simp]: "nat (of_nat n) = n" haftmann@25919: by (simp add: nat int_def) haftmann@25919: haftmann@25919: lemma nat_zero [simp]: "nat 0 = 0" haftmann@25919: by (simp add: Zero_int_def nat) haftmann@25919: haftmann@25919: lemma int_nat_eq [simp]: "of_nat (nat z) = (if 0 \ z then z else 0)" haftmann@25919: by (cases z, simp add: nat le int_def Zero_int_def) haftmann@25919: haftmann@25919: corollary nat_0_le: "0 \ z ==> of_nat (nat z) = z" haftmann@25919: by simp haftmann@25919: haftmann@25919: lemma nat_le_0 [simp]: "z \ 0 ==> nat z = 0" haftmann@25919: by (cases z, simp add: nat le Zero_int_def) haftmann@25919: haftmann@25919: lemma nat_le_eq_zle: "0 < w | 0 \ z ==> (nat w \ nat z) = (w\z)" haftmann@25919: apply (cases w, cases z) haftmann@25919: apply (simp add: nat le linorder_not_le [symmetric] Zero_int_def, arith) haftmann@25919: done haftmann@25919: haftmann@25919: text{*An alternative condition is @{term "0 \ w"} *} haftmann@25919: corollary nat_mono_iff: "0 < z ==> (nat w < nat z) = (w < z)" haftmann@25919: by (simp add: nat_le_eq_zle linorder_not_le [symmetric]) haftmann@25919: haftmann@25919: corollary nat_less_eq_zless: "0 \ w ==> (nat w < nat z) = (w z" and "\m. z = of_nat m \ P" haftmann@25919: shows P haftmann@25919: using assms by (blast dest: nat_0_le sym) haftmann@25919: haftmann@25919: lemma nat_eq_iff: "(nat w = m) = (if 0 \ w then w = of_nat m else m=0)" haftmann@25919: by (cases w, simp add: nat le int_def Zero_int_def, arith) haftmann@25919: haftmann@25919: corollary nat_eq_iff2: "(m = nat w) = (if 0 \ w then w = of_nat m else m=0)" haftmann@25919: by (simp only: eq_commute [of m] nat_eq_iff) haftmann@25919: haftmann@25919: lemma nat_less_iff: "0 \ w ==> (nat w < m) = (w < of_nat m)" haftmann@25919: apply (cases w) haftmann@25919: apply (simp add: nat le int_def Zero_int_def linorder_not_le [symmetric], arith) haftmann@25919: done haftmann@25919: haftmann@25919: lemma int_eq_iff: "(of_nat m = z) = (m = nat z & 0 \ z)" haftmann@25919: by (auto simp add: nat_eq_iff2) haftmann@25919: haftmann@25919: lemma zero_less_nat_eq [simp]: "(0 < nat z) = (0 < z)" haftmann@25919: by (insert zless_nat_conj [of 0], auto) haftmann@25919: haftmann@25919: lemma nat_add_distrib: haftmann@25919: "[| (0::int) \ z; 0 \ z' |] ==> nat (z+z') = nat z + nat z'" haftmann@25919: by (cases z, cases z', simp add: nat add le Zero_int_def) haftmann@25919: haftmann@25919: lemma nat_diff_distrib: haftmann@25919: "[| (0::int) \ z'; z' \ z |] ==> nat (z-z') = nat z - nat z'" haftmann@25919: by (cases z, cases z', haftmann@25919: simp add: nat add minus diff_minus le Zero_int_def) haftmann@25919: haftmann@25919: lemma nat_zminus_int [simp]: "nat (- (of_nat n)) = 0" haftmann@25919: by (simp add: int_def minus nat Zero_int_def) haftmann@25919: haftmann@25919: lemma zless_nat_eq_int_zless: "(m < nat z) = (of_nat m < z)" haftmann@25919: by (cases z, simp add: nat less int_def, arith) haftmann@25919: haftmann@25919: context ring_1 haftmann@25919: begin haftmann@25919: haftmann@25919: lemma of_nat_nat: "0 \ z \ of_nat (nat z) = of_int z" haftmann@25919: by (cases z rule: eq_Abs_Integ) haftmann@25919: (simp add: nat le of_int Zero_int_def of_nat_diff) haftmann@25919: haftmann@25919: end haftmann@25919: haftmann@25919: haftmann@25919: subsection{*Lemmas about the Function @{term of_nat} and Orderings*} haftmann@25919: haftmann@25919: lemma negative_zless_0: "- (of_nat (Suc n)) < (0 \ int)" haftmann@25919: by (simp add: order_less_le del: of_nat_Suc) haftmann@25919: haftmann@25919: lemma negative_zless [iff]: "- (of_nat (Suc n)) < (of_nat m \ int)" haftmann@25919: by (rule negative_zless_0 [THEN order_less_le_trans], simp) haftmann@25919: haftmann@25919: lemma negative_zle_0: "- of_nat n \ (0 \ int)" haftmann@25919: by (simp add: minus_le_iff) haftmann@25919: haftmann@25919: lemma negative_zle [iff]: "- of_nat n \ (of_nat m \ int)" haftmann@25919: by (rule order_trans [OF negative_zle_0 of_nat_0_le_iff]) haftmann@25919: haftmann@25919: lemma not_zle_0_negative [simp]: "~ (0 \ - (of_nat (Suc n) \ int))" haftmann@25919: by (subst le_minus_iff, simp del: of_nat_Suc) haftmann@25919: haftmann@25919: lemma int_zle_neg: "((of_nat n \ int) \ - of_nat m) = (n = 0 & m = 0)" haftmann@25919: by (simp add: int_def le minus Zero_int_def) haftmann@25919: haftmann@25919: lemma not_int_zless_negative [simp]: "~ ((of_nat n \ int) < - of_nat m)" haftmann@25919: by (simp add: linorder_not_less) haftmann@25919: haftmann@25919: lemma negative_eq_positive [simp]: "((- of_nat n \ int) = of_nat m) = (n = 0 & m = 0)" haftmann@25919: by (force simp add: order_eq_iff [of "- of_nat n"] int_zle_neg) haftmann@25919: haftmann@25919: lemma zle_iff_zadd: "(w\int) \ z \ (\n. z = w + of_nat n)" haftmann@25919: proof - haftmann@25919: have "(w \ z) = (0 \ z - w)" haftmann@25919: by (simp only: le_diff_eq add_0_left) haftmann@25919: also have "\ = (\n. z - w = of_nat n)" haftmann@25919: by (auto elim: zero_le_imp_eq_int) haftmann@25919: also have "\ = (\n. z = w + of_nat n)" haftmann@25919: by (simp only: group_simps) haftmann@25919: finally show ?thesis . haftmann@25919: qed haftmann@25919: haftmann@25919: lemma zadd_int_left: "of_nat m + (of_nat n + z) = of_nat (m + n) + (z\int)" haftmann@25919: by simp haftmann@25919: haftmann@25919: lemma int_Suc0_eq_1: "of_nat (Suc 0) = (1\int)" haftmann@25919: by simp haftmann@25919: haftmann@25919: text{*This version is proved for all ordered rings, not just integers! haftmann@25919: It is proved here because attribute @{text arith_split} is not available haftmann@25919: in theory @{text Ring_and_Field}. haftmann@25919: But is it really better than just rewriting with @{text abs_if}?*} haftmann@25919: lemma abs_split [arith_split,noatp]: haftmann@25919: "P(abs(a::'a::ordered_idom)) = ((0 \ a --> P a) & (a < 0 --> P(-a)))" haftmann@25919: by (force dest: order_less_le_trans simp add: abs_if linorder_not_less) haftmann@25919: haftmann@25919: lemma negD: "(x \ int) < 0 \ \n. x = - (of_nat (Suc n))" haftmann@25919: apply (cases x) haftmann@25919: apply (auto simp add: le minus Zero_int_def int_def order_less_le) haftmann@25919: apply (rule_tac x="y - Suc x" in exI, arith) haftmann@25919: done haftmann@25919: haftmann@25919: haftmann@25919: subsection {* Cases and induction *} haftmann@25919: haftmann@25919: text{*Now we replace the case analysis rule by a more conventional one: haftmann@25919: whether an integer is negative or not.*} haftmann@25919: haftmann@25919: theorem int_cases [cases type: int, case_names nonneg neg]: haftmann@25919: "[|!! n. (z \ int) = of_nat n ==> P; !! n. z = - (of_nat (Suc n)) ==> P |] ==> P" haftmann@25919: apply (cases "z < 0", blast dest!: negD) haftmann@25919: apply (simp add: linorder_not_less del: of_nat_Suc) haftmann@25919: apply auto haftmann@25919: apply (blast dest: nat_0_le [THEN sym]) haftmann@25919: done haftmann@25919: haftmann@25919: theorem int_induct [induct type: int, case_names nonneg neg]: haftmann@25919: "[|!! n. P (of_nat n \ int); !!n. P (- (of_nat (Suc n))) |] ==> P z" haftmann@25919: by (cases z rule: int_cases) auto haftmann@25919: haftmann@25919: text{*Contributed by Brian Huffman*} haftmann@25919: theorem int_diff_cases: haftmann@25919: obtains (diff) m n where "(z\int) = of_nat m - of_nat n" haftmann@25919: apply (cases z rule: eq_Abs_Integ) haftmann@25919: apply (rule_tac m=x and n=y in diff) haftmann@25919: apply (simp add: int_def diff_def minus add) haftmann@25919: done haftmann@25919: haftmann@25919: haftmann@25919: subsection {* Binary representation *} haftmann@25919: haftmann@25919: text {* haftmann@25919: This formalization defines binary arithmetic in terms of the integers haftmann@25919: rather than using a datatype. This avoids multiple representations (leading haftmann@25919: zeroes, etc.) See @{text "ZF/Tools/twos-compl.ML"}, function @{text haftmann@25919: int_of_binary}, for the numerical interpretation. haftmann@25919: haftmann@25919: The representation expects that @{text "(m mod 2)"} is 0 or 1, haftmann@25919: even if m is negative; haftmann@25919: For instance, @{text "-5 div 2 = -3"} and @{text "-5 mod 2 = 1"}; thus haftmann@25919: @{text "-5 = (-3)*2 + 1"}. haftmann@25919: haftmann@25919: This two's complement binary representation derives from the paper haftmann@25919: "An Efficient Representation of Arithmetic for Term Rewriting" by haftmann@25919: Dave Cohen and Phil Watson, Rewriting Techniques and Applications, haftmann@25919: Springer LNCS 488 (240-251), 1991. haftmann@25919: *} haftmann@25919: huffman@28958: subsubsection {* The constructors @{term Bit0}, @{term Bit1}, @{term Pls} and @{term Min} *} huffman@28958: haftmann@25919: definition haftmann@25919: Pls :: int where haftmann@28562: [code del]: "Pls = 0" haftmann@25919: haftmann@25919: definition haftmann@25919: Min :: int where haftmann@28562: [code del]: "Min = - 1" haftmann@25919: haftmann@25919: definition huffman@26086: Bit0 :: "int \ int" where haftmann@28562: [code del]: "Bit0 k = k + k" huffman@26086: huffman@26086: definition huffman@26086: Bit1 :: "int \ int" where haftmann@28562: [code del]: "Bit1 k = 1 + k + k" haftmann@25919: haftmann@25919: class number = type + -- {* for numeric types: nat, int, real, \dots *} haftmann@25919: fixes number_of :: "int \ 'a" haftmann@25919: haftmann@25919: use "Tools/numeral.ML" haftmann@25919: haftmann@25919: syntax haftmann@25919: "_Numeral" :: "num_const \ 'a" ("_") haftmann@25919: haftmann@25919: use "Tools/numeral_syntax.ML" haftmann@25919: setup NumeralSyntax.setup haftmann@25919: haftmann@25919: abbreviation haftmann@25919: "Numeral0 \ number_of Pls" haftmann@25919: haftmann@25919: abbreviation huffman@26086: "Numeral1 \ number_of (Bit1 Pls)" haftmann@25919: haftmann@25919: lemma Let_number_of [simp]: "Let (number_of v) f = f (number_of v)" haftmann@25919: -- {* Unfold all @{text let}s involving constants *} haftmann@25919: unfolding Let_def .. haftmann@25919: haftmann@25919: definition haftmann@25919: succ :: "int \ int" where haftmann@28562: [code del]: "succ k = k + 1" haftmann@25919: haftmann@25919: definition haftmann@25919: pred :: "int \ int" where haftmann@28562: [code del]: "pred k = k - 1" haftmann@25919: haftmann@25919: lemmas haftmann@25919: max_number_of [simp] = max_def haftmann@25919: [of "number_of u" "number_of v", standard, simp] haftmann@25919: and haftmann@25919: min_number_of [simp] = min_def haftmann@25919: [of "number_of u" "number_of v", standard, simp] haftmann@25919: -- {* unfolding @{text minx} and @{text max} on numerals *} haftmann@25919: haftmann@25919: lemmas numeral_simps = huffman@26086: succ_def pred_def Pls_def Min_def Bit0_def Bit1_def haftmann@25919: haftmann@25919: text {* Removal of leading zeroes *} haftmann@25919: huffman@26086: lemma Bit0_Pls [simp, code post]: huffman@26086: "Bit0 Pls = Pls" haftmann@25919: unfolding numeral_simps by simp haftmann@25919: huffman@26086: lemma Bit1_Min [simp, code post]: huffman@26086: "Bit1 Min = Min" haftmann@25919: unfolding numeral_simps by simp haftmann@25919: huffman@26075: lemmas normalize_bin_simps = huffman@26086: Bit0_Pls Bit1_Min huffman@26075: haftmann@25919: huffman@28958: subsubsection {* Successor and predecessor functions *} huffman@28958: huffman@28958: text {* Successor *} huffman@28958: huffman@28958: lemma succ_Pls: huffman@26086: "succ Pls = Bit1 Pls" haftmann@25919: unfolding numeral_simps by simp haftmann@25919: huffman@28958: lemma succ_Min: haftmann@25919: "succ Min = Pls" haftmann@25919: unfolding numeral_simps by simp haftmann@25919: huffman@28958: lemma succ_Bit0: huffman@26086: "succ (Bit0 k) = Bit1 k" haftmann@25919: unfolding numeral_simps by simp haftmann@25919: huffman@28958: lemma succ_Bit1: huffman@26086: "succ (Bit1 k) = Bit0 (succ k)" haftmann@25919: unfolding numeral_simps by simp haftmann@25919: huffman@28958: lemmas succ_bin_simps [simp] = huffman@26086: succ_Pls succ_Min succ_Bit0 succ_Bit1 huffman@26075: huffman@28958: text {* Predecessor *} huffman@28958: huffman@28958: lemma pred_Pls: haftmann@25919: "pred Pls = Min" haftmann@25919: unfolding numeral_simps by simp haftmann@25919: huffman@28958: lemma pred_Min: huffman@26086: "pred Min = Bit0 Min" haftmann@25919: unfolding numeral_simps by simp haftmann@25919: huffman@28958: lemma pred_Bit0: huffman@26086: "pred (Bit0 k) = Bit1 (pred k)" haftmann@25919: unfolding numeral_simps by simp haftmann@25919: huffman@28958: lemma pred_Bit1: huffman@26086: "pred (Bit1 k) = Bit0 k" huffman@26086: unfolding numeral_simps by simp huffman@26086: huffman@28958: lemmas pred_bin_simps [simp] = huffman@26086: pred_Pls pred_Min pred_Bit0 pred_Bit1 huffman@26075: huffman@28958: huffman@28958: subsubsection {* Binary arithmetic *} huffman@28958: huffman@28958: text {* Addition *} huffman@28958: huffman@28958: lemma add_Pls: huffman@28958: "Pls + k = k" huffman@28958: unfolding numeral_simps by simp huffman@28958: huffman@28958: lemma add_Min: huffman@28958: "Min + k = pred k" huffman@28958: unfolding numeral_simps by simp huffman@28958: huffman@28958: lemma add_Bit0_Bit0: huffman@28958: "(Bit0 k) + (Bit0 l) = Bit0 (k + l)" huffman@28958: unfolding numeral_simps by simp huffman@28958: huffman@28958: lemma add_Bit0_Bit1: huffman@28958: "(Bit0 k) + (Bit1 l) = Bit1 (k + l)" huffman@28958: unfolding numeral_simps by simp huffman@28958: huffman@28958: lemma add_Bit1_Bit0: huffman@28958: "(Bit1 k) + (Bit0 l) = Bit1 (k + l)" huffman@28958: unfolding numeral_simps by simp huffman@28958: huffman@28958: lemma add_Bit1_Bit1: huffman@28958: "(Bit1 k) + (Bit1 l) = Bit0 (k + succ l)" huffman@28958: unfolding numeral_simps by simp huffman@28958: huffman@28958: lemma add_Pls_right: huffman@28958: "k + Pls = k" huffman@28958: unfolding numeral_simps by simp huffman@28958: huffman@28958: lemma add_Min_right: huffman@28958: "k + Min = pred k" huffman@28958: unfolding numeral_simps by simp huffman@28958: huffman@28958: lemmas add_bin_simps [simp] = huffman@28958: add_Pls add_Min add_Pls_right add_Min_right huffman@28958: add_Bit0_Bit0 add_Bit0_Bit1 add_Bit1_Bit0 add_Bit1_Bit1 huffman@28958: huffman@28958: text {* Negation *} huffman@28958: huffman@28958: lemma minus_Pls: haftmann@25919: "- Pls = Pls" huffman@28958: unfolding numeral_simps by simp huffman@28958: huffman@28958: lemma minus_Min: huffman@26086: "- Min = Bit1 Pls" huffman@28958: unfolding numeral_simps by simp huffman@28958: huffman@28958: lemma minus_Bit0: huffman@26086: "- (Bit0 k) = Bit0 (- k)" huffman@28958: unfolding numeral_simps by simp huffman@28958: huffman@28958: lemma minus_Bit1: huffman@26086: "- (Bit1 k) = Bit1 (pred (- k))" huffman@26086: unfolding numeral_simps by simp haftmann@25919: huffman@28958: lemmas minus_bin_simps [simp] = huffman@26086: minus_Pls minus_Min minus_Bit0 minus_Bit1 huffman@26075: huffman@28958: text {* Subtraction *} huffman@28958: huffman@28958: lemma diff_Pls: huffman@28958: "Pls - k = - k" huffman@28958: unfolding numeral_simps by simp huffman@28958: huffman@28958: lemma diff_Min: huffman@28958: "Min - k = pred (- k)" huffman@28958: unfolding numeral_simps by simp huffman@28958: huffman@28958: lemma diff_Bit0_Bit0: huffman@28958: "(Bit0 k) - (Bit0 l) = Bit0 (k - l)" huffman@28958: unfolding numeral_simps by simp huffman@28958: huffman@28958: lemma diff_Bit0_Bit1: huffman@28958: "(Bit0 k) - (Bit1 l) = Bit1 (pred k - l)" huffman@28958: unfolding numeral_simps by simp huffman@28958: huffman@28958: lemma diff_Bit1_Bit0: huffman@28958: "(Bit1 k) - (Bit0 l) = Bit1 (k - l)" haftmann@25919: unfolding numeral_simps by simp haftmann@25919: huffman@28958: lemma diff_Bit1_Bit1: huffman@28958: "(Bit1 k) - (Bit1 l) = Bit0 (k - l)" haftmann@25919: unfolding numeral_simps by simp haftmann@25919: huffman@28958: lemma diff_Pls_right: huffman@28958: "k - Pls = k" huffman@28958: unfolding numeral_simps by simp huffman@28958: huffman@28958: lemma diff_Min_right: huffman@28958: "k - Min = succ k" haftmann@25919: unfolding numeral_simps by simp haftmann@25919: huffman@28958: lemmas diff_bin_simps [simp] = huffman@28958: diff_Pls diff_Min diff_Pls_right diff_Min_right huffman@28958: diff_Bit0_Bit0 diff_Bit0_Bit1 diff_Bit1_Bit0 diff_Bit1_Bit1 huffman@28958: huffman@28958: text {* Multiplication *} huffman@28958: huffman@28958: lemma mult_Pls: huffman@28958: "Pls * w = Pls" huffman@26086: unfolding numeral_simps by simp haftmann@25919: huffman@28958: lemma mult_Min: haftmann@25919: "Min * k = - k" haftmann@25919: unfolding numeral_simps by simp haftmann@25919: huffman@28958: lemma mult_Bit0: huffman@26086: "(Bit0 k) * l = Bit0 (k * l)" huffman@26086: unfolding numeral_simps int_distrib by simp haftmann@25919: huffman@28958: lemma mult_Bit1: huffman@26086: "(Bit1 k) * l = (Bit0 (k * l)) + l" huffman@28958: unfolding numeral_simps int_distrib by simp huffman@28958: huffman@28958: lemmas mult_bin_simps [simp] = huffman@26086: mult_Pls mult_Min mult_Bit0 mult_Bit1 huffman@26075: haftmann@25919: huffman@28958: subsubsection {* Binary comparisons *} huffman@28958: huffman@28958: text {* Preliminaries *} huffman@28958: huffman@28958: lemma even_less_0_iff: huffman@28958: "a + a < 0 \ a < (0::'a::ordered_idom)" huffman@28958: proof - huffman@28958: have "a + a < 0 \ (1+1)*a < 0" by (simp add: left_distrib) huffman@28958: also have "(1+1)*a < 0 \ a < 0" huffman@28958: by (simp add: mult_less_0_iff zero_less_two huffman@28958: order_less_not_sym [OF zero_less_two]) huffman@28958: finally show ?thesis . huffman@28958: qed huffman@28958: huffman@28958: lemma le_imp_0_less: huffman@28958: assumes le: "0 \ z" huffman@28958: shows "(0::int) < 1 + z" huffman@28958: proof - huffman@28958: have "0 \ z" by fact huffman@28958: also have "... < z + 1" by (rule less_add_one) huffman@28958: also have "... = 1 + z" by (simp add: add_ac) huffman@28958: finally show "0 < 1 + z" . huffman@28958: qed huffman@28958: huffman@28958: lemma odd_less_0_iff: huffman@28958: "(1 + z + z < 0) = (z < (0::int))" huffman@28958: proof (cases z rule: int_cases) huffman@28958: case (nonneg n) huffman@28958: thus ?thesis by (simp add: linorder_not_less add_assoc add_increasing huffman@28958: le_imp_0_less [THEN order_less_imp_le]) huffman@28958: next huffman@28958: case (neg n) huffman@28958: thus ?thesis by (simp del: of_nat_Suc of_nat_add huffman@28958: add: compare_rls of_nat_1 [symmetric] of_nat_add [symmetric]) huffman@28958: qed huffman@28958: huffman@28958: lemma neg_simps: huffman@28958: "Pls < 0 \ False" huffman@28958: "Min < 0 \ True" huffman@28958: "Bit0 w < 0 \ w < 0" huffman@28958: "Bit1 w < 0 \ w < 0" huffman@28958: unfolding numeral_simps huffman@28958: by (simp_all add: even_less_0_iff odd_less_0_iff) huffman@28958: huffman@28958: lemma less_bin_lemma: "k < l \ k - l < (0::int)" huffman@28958: by simp huffman@28958: huffman@28958: lemma le_iff_pred_less: "k \ l \ pred k < l" huffman@28958: unfolding numeral_simps huffman@28958: proof huffman@28958: have "k - 1 < k" by simp huffman@28958: also assume "k \ l" huffman@28958: finally show "k - 1 < l" . huffman@28958: next huffman@28958: assume "k - 1 < l" huffman@28958: hence "(k - 1) + 1 \ l" by (rule zless_imp_add1_zle) huffman@28958: thus "k \ l" by simp huffman@28958: qed huffman@28958: huffman@28958: lemma succ_pred: "succ (pred x) = x" huffman@28958: unfolding numeral_simps by simp huffman@28958: huffman@28958: text {* Less-than *} huffman@28958: huffman@28958: lemma less_bin_simps [simp]: huffman@28958: "Pls < Pls \ False" huffman@28958: "Pls < Min \ False" huffman@28958: "Pls < Bit0 k \ Pls < k" huffman@28958: "Pls < Bit1 k \ Pls \ k" huffman@28958: "Min < Pls \ True" huffman@28958: "Min < Min \ False" huffman@28958: "Min < Bit0 k \ Min < k" huffman@28958: "Min < Bit1 k \ Min < k" huffman@28958: "Bit0 k < Pls \ k < Pls" huffman@28958: "Bit0 k < Min \ k \ Min" huffman@28958: "Bit1 k < Pls \ k < Pls" huffman@28958: "Bit1 k < Min \ k < Min" huffman@28958: "Bit0 k < Bit0 l \ k < l" huffman@28958: "Bit0 k < Bit1 l \ k \ l" huffman@28958: "Bit1 k < Bit0 l \ k < l" huffman@28958: "Bit1 k < Bit1 l \ k < l" huffman@28958: unfolding le_iff_pred_less huffman@28958: less_bin_lemma [of Pls] huffman@28958: less_bin_lemma [of Min] huffman@28958: less_bin_lemma [of "k"] huffman@28958: less_bin_lemma [of "Bit0 k"] huffman@28958: less_bin_lemma [of "Bit1 k"] huffman@28958: less_bin_lemma [of "pred Pls"] huffman@28958: less_bin_lemma [of "pred k"] huffman@28958: by (simp_all add: neg_simps succ_pred) huffman@28958: huffman@28958: text {* Less-than-or-equal *} huffman@28958: huffman@28958: lemma le_bin_simps [simp]: huffman@28958: "Pls \ Pls \ True" huffman@28958: "Pls \ Min \ False" huffman@28958: "Pls \ Bit0 k \ Pls \ k" huffman@28958: "Pls \ Bit1 k \ Pls \ k" huffman@28958: "Min \ Pls \ True" huffman@28958: "Min \ Min \ True" huffman@28958: "Min \ Bit0 k \ Min < k" huffman@28958: "Min \ Bit1 k \ Min \ k" huffman@28958: "Bit0 k \ Pls \ k \ Pls" huffman@28958: "Bit0 k \ Min \ k \ Min" huffman@28958: "Bit1 k \ Pls \ k < Pls" huffman@28958: "Bit1 k \ Min \ k \ Min" huffman@28958: "Bit0 k \ Bit0 l \ k \ l" huffman@28958: "Bit0 k \ Bit1 l \ k \ l" huffman@28958: "Bit1 k \ Bit0 l \ k < l" huffman@28958: "Bit1 k \ Bit1 l \ k \ l" huffman@28958: unfolding not_less [symmetric] huffman@28958: by (simp_all add: not_le) huffman@28958: huffman@28958: text {* Equality *} huffman@28958: huffman@28958: lemma eq_bin_simps [simp]: huffman@28958: "Pls = Pls \ True" huffman@28958: "Pls = Min \ False" huffman@28958: "Pls = Bit0 l \ Pls = l" huffman@28958: "Pls = Bit1 l \ False" huffman@28958: "Min = Pls \ False" huffman@28958: "Min = Min \ True" huffman@28958: "Min = Bit0 l \ False" huffman@28958: "Min = Bit1 l \ Min = l" huffman@28958: "Bit0 k = Pls \ k = Pls" huffman@28958: "Bit0 k = Min \ False" huffman@28958: "Bit1 k = Pls \ False" huffman@28958: "Bit1 k = Min \ k = Min" huffman@28958: "Bit0 k = Bit0 l \ k = l" huffman@28958: "Bit0 k = Bit1 l \ False" huffman@28958: "Bit1 k = Bit0 l \ False" huffman@28958: "Bit1 k = Bit1 l \ k = l" huffman@28958: unfolding order_eq_iff [where 'a=int] huffman@28958: by (simp_all add: not_less) huffman@28958: huffman@28958: haftmann@25919: subsection {* Converting Numerals to Rings: @{term number_of} *} haftmann@25919: haftmann@25919: class number_ring = number + comm_ring_1 + haftmann@25919: assumes number_of_eq: "number_of k = of_int k" haftmann@25919: haftmann@25919: text {* self-embedding of the integers *} haftmann@25919: haftmann@25919: instantiation int :: number_ring haftmann@25919: begin haftmann@25919: haftmann@28724: definition int_number_of_def [code del]: haftmann@28724: "number_of w = (of_int w \ int)" haftmann@25919: haftmann@28724: instance proof haftmann@28724: qed (simp only: int_number_of_def) haftmann@25919: haftmann@25919: end haftmann@25919: haftmann@25919: lemma number_of_is_id: haftmann@25919: "number_of (k::int) = k" haftmann@25919: unfolding int_number_of_def by simp haftmann@25919: haftmann@25919: lemma number_of_succ: haftmann@25919: "number_of (succ k) = (1 + number_of k ::'a::number_ring)" haftmann@25919: unfolding number_of_eq numeral_simps by simp haftmann@25919: haftmann@25919: lemma number_of_pred: haftmann@25919: "number_of (pred w) = (- 1 + number_of w ::'a::number_ring)" haftmann@25919: unfolding number_of_eq numeral_simps by simp haftmann@25919: haftmann@25919: lemma number_of_minus: haftmann@25919: "number_of (uminus w) = (- (number_of w)::'a::number_ring)" huffman@28958: unfolding number_of_eq by (rule of_int_minus) haftmann@25919: haftmann@25919: lemma number_of_add: haftmann@25919: "number_of (v + w) = (number_of v + number_of w::'a::number_ring)" huffman@28958: unfolding number_of_eq by (rule of_int_add) huffman@28958: huffman@28958: lemma number_of_diff: huffman@28958: "number_of (v - w) = (number_of v - number_of w::'a::number_ring)" huffman@28958: unfolding number_of_eq by (rule of_int_diff) haftmann@25919: haftmann@25919: lemma number_of_mult: haftmann@25919: "number_of (v * w) = (number_of v * number_of w::'a::number_ring)" huffman@28958: unfolding number_of_eq by (rule of_int_mult) haftmann@25919: haftmann@25919: text {* haftmann@25919: The correctness of shifting. haftmann@25919: But it doesn't seem to give a measurable speed-up. haftmann@25919: *} haftmann@25919: huffman@26086: lemma double_number_of_Bit0: huffman@26086: "(1 + 1) * number_of w = (number_of (Bit0 w) ::'a::number_ring)" haftmann@25919: unfolding number_of_eq numeral_simps left_distrib by simp haftmann@25919: haftmann@25919: text {* haftmann@25919: Converting numerals 0 and 1 to their abstract versions. haftmann@25919: *} haftmann@25919: haftmann@25919: lemma numeral_0_eq_0 [simp]: haftmann@25919: "Numeral0 = (0::'a::number_ring)" haftmann@25919: unfolding number_of_eq numeral_simps by simp haftmann@25919: haftmann@25919: lemma numeral_1_eq_1 [simp]: haftmann@25919: "Numeral1 = (1::'a::number_ring)" haftmann@25919: unfolding number_of_eq numeral_simps by simp haftmann@25919: haftmann@25919: text {* haftmann@25919: Special-case simplification for small constants. haftmann@25919: *} haftmann@25919: haftmann@25919: text{* haftmann@25919: Unary minus for the abstract constant 1. Cannot be inserted haftmann@25919: as a simprule until later: it is @{text number_of_Min} re-oriented! haftmann@25919: *} haftmann@25919: haftmann@25919: lemma numeral_m1_eq_minus_1: haftmann@25919: "(-1::'a::number_ring) = - 1" haftmann@25919: unfolding number_of_eq numeral_simps by simp haftmann@25919: haftmann@25919: lemma mult_minus1 [simp]: haftmann@25919: "-1 * z = -(z::'a::number_ring)" haftmann@25919: unfolding number_of_eq numeral_simps by simp haftmann@25919: haftmann@25919: lemma mult_minus1_right [simp]: haftmann@25919: "z * -1 = -(z::'a::number_ring)" haftmann@25919: unfolding number_of_eq numeral_simps by simp haftmann@25919: haftmann@25919: (*Negation of a coefficient*) haftmann@25919: lemma minus_number_of_mult [simp]: haftmann@25919: "- (number_of w) * z = number_of (uminus w) * (z::'a::number_ring)" haftmann@25919: unfolding number_of_eq by simp haftmann@25919: haftmann@25919: text {* Subtraction *} haftmann@25919: haftmann@25919: lemma diff_number_of_eq: haftmann@25919: "number_of v - number_of w = haftmann@25919: (number_of (v + uminus w)::'a::number_ring)" haftmann@25919: unfolding number_of_eq by simp haftmann@25919: haftmann@25919: lemma number_of_Pls: haftmann@25919: "number_of Pls = (0::'a::number_ring)" haftmann@25919: unfolding number_of_eq numeral_simps by simp haftmann@25919: haftmann@25919: lemma number_of_Min: haftmann@25919: "number_of Min = (- 1::'a::number_ring)" haftmann@25919: unfolding number_of_eq numeral_simps by simp haftmann@25919: huffman@26086: lemma number_of_Bit0: huffman@26086: "number_of (Bit0 w) = (0::'a::number_ring) + (number_of w) + (number_of w)" huffman@26086: unfolding number_of_eq numeral_simps by simp huffman@26086: huffman@26086: lemma number_of_Bit1: huffman@26086: "number_of (Bit1 w) = (1::'a::number_ring) + (number_of w) + (number_of w)" huffman@26086: unfolding number_of_eq numeral_simps by simp haftmann@25919: haftmann@25919: huffman@28958: subsubsection {* Equality of Binary Numbers *} haftmann@25919: haftmann@25919: text {* First version by Norbert Voelker *} haftmann@25919: haftmann@25919: definition haftmann@25919: neg :: "'a\ordered_idom \ bool" haftmann@25919: where haftmann@25919: "neg Z \ Z < 0" haftmann@25919: haftmann@25919: definition (*for simplifying equalities*) haftmann@25919: iszero :: "'a\semiring_1 \ bool" haftmann@25919: where haftmann@25919: "iszero z \ z = 0" haftmann@25919: haftmann@25919: lemma not_neg_int [simp]: "~ neg (of_nat n)" haftmann@25919: by (simp add: neg_def) haftmann@25919: haftmann@25919: lemma neg_zminus_int [simp]: "neg (- (of_nat (Suc n)))" haftmann@25919: by (simp add: neg_def neg_less_0_iff_less del: of_nat_Suc) haftmann@25919: haftmann@25919: lemmas neg_eq_less_0 = neg_def haftmann@25919: haftmann@25919: lemma not_neg_eq_ge_0: "(~neg x) = (0 \ x)" haftmann@25919: by (simp add: neg_def linorder_not_less) haftmann@25919: haftmann@25919: text{*To simplify inequalities when Numeral1 can get simplified to 1*} haftmann@25919: haftmann@25919: lemma not_neg_0: "~ neg 0" haftmann@25919: by (simp add: One_int_def neg_def) haftmann@25919: haftmann@25919: lemma not_neg_1: "~ neg 1" haftmann@25919: by (simp add: neg_def linorder_not_less zero_le_one) haftmann@25919: haftmann@25919: lemma iszero_0: "iszero 0" haftmann@25919: by (simp add: iszero_def) haftmann@25919: haftmann@25919: lemma not_iszero_1: "~ iszero 1" haftmann@25919: by (simp add: iszero_def eq_commute) haftmann@25919: haftmann@25919: lemma neg_nat: "neg z ==> nat z = 0" haftmann@25919: by (simp add: neg_def order_less_imp_le) haftmann@25919: haftmann@25919: lemma not_neg_nat: "~ neg z ==> of_nat (nat z) = z" haftmann@25919: by (simp add: linorder_not_less neg_def) haftmann@25919: haftmann@25919: lemma eq_number_of_eq: haftmann@25919: "((number_of x::'a::number_ring) = number_of y) = haftmann@25919: iszero (number_of (x + uminus y) :: 'a)" haftmann@25919: unfolding iszero_def number_of_add number_of_minus haftmann@25919: by (simp add: compare_rls) haftmann@25919: haftmann@25919: lemma iszero_number_of_Pls: haftmann@25919: "iszero ((number_of Pls)::'a::number_ring)" haftmann@25919: unfolding iszero_def numeral_0_eq_0 .. haftmann@25919: haftmann@25919: lemma nonzero_number_of_Min: haftmann@25919: "~ iszero ((number_of Min)::'a::number_ring)" haftmann@25919: unfolding iszero_def numeral_m1_eq_minus_1 by simp haftmann@25919: haftmann@25919: huffman@28958: subsubsection {* Comparisons, for Ordered Rings *} haftmann@25919: haftmann@25919: lemmas double_eq_0_iff = double_zero haftmann@25919: haftmann@25919: lemma odd_nonzero: haftmann@25919: "1 + z + z \ (0::int)"; haftmann@25919: proof (cases z rule: int_cases) haftmann@25919: case (nonneg n) haftmann@25919: have le: "0 \ z+z" by (simp add: nonneg add_increasing) haftmann@25919: thus ?thesis using le_imp_0_less [OF le] haftmann@25919: by (auto simp add: add_assoc) haftmann@25919: next haftmann@25919: case (neg n) haftmann@25919: show ?thesis haftmann@25919: proof haftmann@25919: assume eq: "1 + z + z = 0" haftmann@25919: have "(0::int) < 1 + (of_nat n + of_nat n)" haftmann@25919: by (simp add: le_imp_0_less add_increasing) haftmann@25919: also have "... = - (1 + z + z)" haftmann@25919: by (simp add: neg add_assoc [symmetric]) haftmann@25919: also have "... = 0" by (simp add: eq) haftmann@25919: finally have "0<0" .. haftmann@25919: thus False by blast haftmann@25919: qed haftmann@25919: qed haftmann@25919: huffman@26086: lemma iszero_number_of_Bit0: huffman@26086: "iszero (number_of (Bit0 w)::'a) = huffman@26086: iszero (number_of w::'a::{ring_char_0,number_ring})" haftmann@25919: proof - haftmann@25919: have "(of_int w + of_int w = (0::'a)) \ (w = 0)" haftmann@25919: proof - haftmann@25919: assume eq: "of_int w + of_int w = (0::'a)" haftmann@25919: then have "of_int (w + w) = (of_int 0 :: 'a)" by simp haftmann@25919: then have "w + w = 0" by (simp only: of_int_eq_iff) haftmann@25919: then show "w = 0" by (simp only: double_eq_0_iff) haftmann@25919: qed huffman@26086: thus ?thesis huffman@26086: by (auto simp add: iszero_def number_of_eq numeral_simps) huffman@26086: qed huffman@26086: huffman@26086: lemma iszero_number_of_Bit1: huffman@26086: "~ iszero (number_of (Bit1 w)::'a::{ring_char_0,number_ring})" huffman@26086: proof - huffman@26086: have "1 + of_int w + of_int w \ (0::'a)" haftmann@25919: proof haftmann@25919: assume eq: "1 + of_int w + of_int w = (0::'a)" haftmann@25919: hence "of_int (1 + w + w) = (of_int 0 :: 'a)" by simp haftmann@25919: hence "1 + w + w = 0" by (simp only: of_int_eq_iff) haftmann@25919: with odd_nonzero show False by blast haftmann@25919: qed huffman@26086: thus ?thesis huffman@26086: by (auto simp add: iszero_def number_of_eq numeral_simps) haftmann@25919: qed haftmann@25919: haftmann@25919: huffman@28958: subsubsection {* The Less-Than Relation *} haftmann@25919: haftmann@25919: lemma less_number_of_eq_neg: haftmann@25919: "((number_of x::'a::{ordered_idom,number_ring}) < number_of y) haftmann@25919: = neg (number_of (x + uminus y) :: 'a)" haftmann@25919: apply (subst less_iff_diff_less_0) haftmann@25919: apply (simp add: neg_def diff_minus number_of_add number_of_minus) haftmann@25919: done haftmann@25919: haftmann@25919: text {* haftmann@25919: If @{term Numeral0} is rewritten to 0 then this rule can't be applied: haftmann@25919: @{term Numeral0} IS @{term "number_of Pls"} haftmann@25919: *} haftmann@25919: haftmann@25919: lemma not_neg_number_of_Pls: haftmann@25919: "~ neg (number_of Pls ::'a::{ordered_idom,number_ring})" haftmann@25919: by (simp add: neg_def numeral_0_eq_0) haftmann@25919: haftmann@25919: lemma neg_number_of_Min: haftmann@25919: "neg (number_of Min ::'a::{ordered_idom,number_ring})" haftmann@25919: by (simp add: neg_def zero_less_one numeral_m1_eq_minus_1) haftmann@25919: haftmann@25919: lemma double_less_0_iff: haftmann@25919: "(a + a < 0) = (a < (0::'a::ordered_idom))" haftmann@25919: proof - haftmann@25919: have "(a + a < 0) = ((1+1)*a < 0)" by (simp add: left_distrib) haftmann@25919: also have "... = (a < 0)" haftmann@25919: by (simp add: mult_less_0_iff zero_less_two haftmann@25919: order_less_not_sym [OF zero_less_two]) haftmann@25919: finally show ?thesis . haftmann@25919: qed haftmann@25919: haftmann@25919: lemma odd_less_0: haftmann@25919: "(1 + z + z < 0) = (z < (0::int))"; haftmann@25919: proof (cases z rule: int_cases) haftmann@25919: case (nonneg n) haftmann@25919: thus ?thesis by (simp add: linorder_not_less add_assoc add_increasing haftmann@25919: le_imp_0_less [THEN order_less_imp_le]) haftmann@25919: next haftmann@25919: case (neg n) haftmann@25919: thus ?thesis by (simp del: of_nat_Suc of_nat_add haftmann@25919: add: compare_rls of_nat_1 [symmetric] of_nat_add [symmetric]) haftmann@25919: qed haftmann@25919: huffman@26086: lemma neg_number_of_Bit0: huffman@26086: "neg (number_of (Bit0 w)::'a) = huffman@26086: neg (number_of w :: 'a::{ordered_idom,number_ring})" huffman@26086: by (simp add: neg_def number_of_eq numeral_simps double_less_0_iff) huffman@26086: huffman@26086: lemma neg_number_of_Bit1: huffman@26086: "neg (number_of (Bit1 w)::'a) = haftmann@25919: neg (number_of w :: 'a::{ordered_idom,number_ring})" haftmann@25919: proof - haftmann@25919: have "((1::'a) + of_int w + of_int w < 0) = (of_int (1 + w + w) < (of_int 0 :: 'a))" haftmann@25919: by simp haftmann@25919: also have "... = (w < 0)" by (simp only: of_int_less_iff odd_less_0) haftmann@25919: finally show ?thesis huffman@26086: by (simp add: neg_def number_of_eq numeral_simps) haftmann@25919: qed haftmann@25919: haftmann@25919: haftmann@25919: text {* Less-Than or Equals *} haftmann@25919: haftmann@25919: text {* Reduces @{term "a\b"} to @{term "~ (b number_of y) haftmann@25919: = (~ (neg (number_of (y + uminus x) :: 'a)))" haftmann@25919: by (simp add: le_number_of_eq_not_less less_number_of_eq_neg) haftmann@25919: haftmann@25919: haftmann@25919: text {* Absolute value (@{term abs}) *} haftmann@25919: haftmann@25919: lemma abs_number_of: haftmann@25919: "abs(number_of x::'a::{ordered_idom,number_ring}) = haftmann@25919: (if number_of x < (0::'a) then -number_of x else number_of x)" haftmann@25919: by (simp add: abs_if) haftmann@25919: haftmann@25919: haftmann@25919: text {* Re-orientation of the equation nnn=x *} haftmann@25919: haftmann@25919: lemma number_of_reorient: haftmann@25919: "(number_of w = x) = (x = number_of w)" haftmann@25919: by auto haftmann@25919: haftmann@25919: huffman@28958: subsubsection {* Simplification of arithmetic operations on integer constants. *} haftmann@25919: haftmann@25919: lemmas arith_extra_simps [standard, simp] = haftmann@25919: number_of_add [symmetric] huffman@28958: number_of_minus [symmetric] huffman@28958: numeral_m1_eq_minus_1 [symmetric] haftmann@25919: number_of_mult [symmetric] haftmann@25919: diff_number_of_eq abs_number_of haftmann@25919: haftmann@25919: text {* haftmann@25919: For making a minimal simpset, one must include these default simprules. haftmann@25919: Also include @{text simp_thms}. haftmann@25919: *} haftmann@25919: haftmann@25919: lemmas arith_simps = huffman@26075: normalize_bin_simps pred_bin_simps succ_bin_simps huffman@26075: add_bin_simps minus_bin_simps mult_bin_simps haftmann@25919: abs_zero abs_one arith_extra_simps haftmann@25919: haftmann@25919: text {* Simplification of relational operations *} haftmann@25919: huffman@28962: lemma less_number_of [simp]: huffman@28962: "(number_of x::'a::{ordered_idom,number_ring}) < number_of y \ x < y" huffman@28962: unfolding number_of_eq by (rule of_int_less_iff) huffman@28962: huffman@28962: lemma le_number_of [simp]: huffman@28962: "(number_of x::'a::{ordered_idom,number_ring}) \ number_of y \ x \ y" huffman@28962: unfolding number_of_eq by (rule of_int_le_iff) huffman@28962: huffman@28967: lemma eq_number_of [simp]: huffman@28967: "(number_of x::'a::{ring_char_0,number_ring}) = number_of y \ x = y" huffman@28967: unfolding number_of_eq by (rule of_int_eq_iff) huffman@28967: haftmann@25919: lemmas rel_simps [simp] = huffman@28962: less_number_of less_bin_simps huffman@28962: le_number_of le_bin_simps huffman@28967: eq_number_of eq_bin_simps huffman@28967: iszero_0 nonzero_number_of_Min huffman@26086: iszero_number_of_Bit0 iszero_number_of_Bit1 haftmann@25919: not_neg_number_of_Pls not_neg_0 not_neg_1 not_iszero_1 huffman@26086: neg_number_of_Min neg_number_of_Bit0 neg_number_of_Bit1 haftmann@25919: (* iszero_number_of_Pls would never be used haftmann@25919: because its lhs simplifies to "iszero 0" *) haftmann@25919: haftmann@25919: huffman@28958: subsubsection {* Simplification of arithmetic when nested to the right. *} haftmann@25919: haftmann@25919: lemma add_number_of_left [simp]: haftmann@25919: "number_of v + (number_of w + z) = haftmann@25919: (number_of(v + w) + z::'a::number_ring)" haftmann@25919: by (simp add: add_assoc [symmetric]) haftmann@25919: haftmann@25919: lemma mult_number_of_left [simp]: haftmann@25919: "number_of v * (number_of w * z) = haftmann@25919: (number_of(v * w) * z::'a::number_ring)" haftmann@25919: by (simp add: mult_assoc [symmetric]) haftmann@25919: haftmann@25919: lemma add_number_of_diff1: haftmann@25919: "number_of v + (number_of w - c) = haftmann@25919: number_of(v + w) - (c::'a::number_ring)" haftmann@25919: by (simp add: diff_minus add_number_of_left) haftmann@25919: haftmann@25919: lemma add_number_of_diff2 [simp]: haftmann@25919: "number_of v + (c - number_of w) = haftmann@25919: number_of (v + uminus w) + (c::'a::number_ring)" haftmann@25919: apply (subst diff_number_of_eq [symmetric]) haftmann@25919: apply (simp only: compare_rls) haftmann@25919: done haftmann@25919: haftmann@25919: haftmann@25919: subsection {* The Set of Integers *} haftmann@25919: haftmann@25919: context ring_1 haftmann@25919: begin haftmann@25919: haftmann@25919: definition haftmann@25919: Ints :: "'a set" haftmann@25919: where haftmann@28562: [code del]: "Ints = range of_int" haftmann@25919: haftmann@25919: end haftmann@25919: haftmann@25919: notation (xsymbols) haftmann@25919: Ints ("\") haftmann@25919: haftmann@25919: context ring_1 haftmann@25919: begin haftmann@25919: haftmann@25919: lemma Ints_0 [simp]: "0 \ \" haftmann@25919: apply (simp add: Ints_def) haftmann@25919: apply (rule range_eqI) haftmann@25919: apply (rule of_int_0 [symmetric]) haftmann@25919: done haftmann@25919: haftmann@25919: lemma Ints_1 [simp]: "1 \ \" haftmann@25919: apply (simp add: Ints_def) haftmann@25919: apply (rule range_eqI) haftmann@25919: apply (rule of_int_1 [symmetric]) haftmann@25919: done haftmann@25919: haftmann@25919: lemma Ints_add [simp]: "a \ \ \ b \ \ \ a + b \ \" haftmann@25919: apply (auto simp add: Ints_def) haftmann@25919: apply (rule range_eqI) haftmann@25919: apply (rule of_int_add [symmetric]) haftmann@25919: done haftmann@25919: haftmann@25919: lemma Ints_minus [simp]: "a \ \ \ -a \ \" haftmann@25919: apply (auto simp add: Ints_def) haftmann@25919: apply (rule range_eqI) haftmann@25919: apply (rule of_int_minus [symmetric]) haftmann@25919: done haftmann@25919: haftmann@25919: lemma Ints_mult [simp]: "a \ \ \ b \ \ \ a * b \ \" haftmann@25919: apply (auto simp add: Ints_def) haftmann@25919: apply (rule range_eqI) haftmann@25919: apply (rule of_int_mult [symmetric]) haftmann@25919: done haftmann@25919: haftmann@25919: lemma Ints_cases [cases set: Ints]: haftmann@25919: assumes "q \ \" haftmann@25919: obtains (of_int) z where "q = of_int z" haftmann@25919: unfolding Ints_def haftmann@25919: proof - haftmann@25919: from `q \ \` have "q \ range of_int" unfolding Ints_def . haftmann@25919: then obtain z where "q = of_int z" .. haftmann@25919: then show thesis .. haftmann@25919: qed haftmann@25919: haftmann@25919: lemma Ints_induct [case_names of_int, induct set: Ints]: haftmann@25919: "q \ \ \ (\z. P (of_int z)) \ P q" haftmann@25919: by (rule Ints_cases) auto haftmann@25919: haftmann@25919: end haftmann@25919: haftmann@25919: lemma Ints_diff [simp]: "a \ \ \ b \ \ \ a-b \ \" haftmann@25919: apply (auto simp add: Ints_def) haftmann@25919: apply (rule range_eqI) haftmann@25919: apply (rule of_int_diff [symmetric]) haftmann@25919: done haftmann@25919: haftmann@25919: text {* The premise involving @{term Ints} prevents @{term "a = 1/2"}. *} haftmann@25919: haftmann@25919: lemma Ints_double_eq_0_iff: haftmann@25919: assumes in_Ints: "a \ Ints" haftmann@25919: shows "(a + a = 0) = (a = (0::'a::ring_char_0))" haftmann@25919: proof - haftmann@25919: from in_Ints have "a \ range of_int" unfolding Ints_def [symmetric] . haftmann@25919: then obtain z where a: "a = of_int z" .. haftmann@25919: show ?thesis haftmann@25919: proof haftmann@25919: assume "a = 0" haftmann@25919: thus "a + a = 0" by simp haftmann@25919: next haftmann@25919: assume eq: "a + a = 0" haftmann@25919: hence "of_int (z + z) = (of_int 0 :: 'a)" by (simp add: a) haftmann@25919: hence "z + z = 0" by (simp only: of_int_eq_iff) haftmann@25919: hence "z = 0" by (simp only: double_eq_0_iff) haftmann@25919: thus "a = 0" by (simp add: a) haftmann@25919: qed haftmann@25919: qed haftmann@25919: haftmann@25919: lemma Ints_odd_nonzero: haftmann@25919: assumes in_Ints: "a \ Ints" haftmann@25919: shows "1 + a + a \ (0::'a::ring_char_0)" haftmann@25919: proof - haftmann@25919: from in_Ints have "a \ range of_int" unfolding Ints_def [symmetric] . haftmann@25919: then obtain z where a: "a = of_int z" .. haftmann@25919: show ?thesis haftmann@25919: proof haftmann@25919: assume eq: "1 + a + a = 0" haftmann@25919: hence "of_int (1 + z + z) = (of_int 0 :: 'a)" by (simp add: a) haftmann@25919: hence "1 + z + z = 0" by (simp only: of_int_eq_iff) haftmann@25919: with odd_nonzero show False by blast haftmann@25919: qed haftmann@25919: qed haftmann@25919: haftmann@25919: lemma Ints_number_of: haftmann@25919: "(number_of w :: 'a::number_ring) \ Ints" haftmann@25919: unfolding number_of_eq Ints_def by simp haftmann@25919: haftmann@25919: lemma Ints_odd_less_0: haftmann@25919: assumes in_Ints: "a \ Ints" haftmann@25919: shows "(1 + a + a < 0) = (a < (0::'a::ordered_idom))"; haftmann@25919: proof - haftmann@25919: from in_Ints have "a \ range of_int" unfolding Ints_def [symmetric] . haftmann@25919: then obtain z where a: "a = of_int z" .. haftmann@25919: hence "((1::'a) + a + a < 0) = (of_int (1 + z + z) < (of_int 0 :: 'a))" haftmann@25919: by (simp add: a) haftmann@25919: also have "... = (z < 0)" by (simp only: of_int_less_iff odd_less_0) haftmann@25919: also have "... = (a < 0)" by (simp add: a) haftmann@25919: finally show ?thesis . haftmann@25919: qed haftmann@25919: haftmann@25919: haftmann@25919: subsection {* @{term setsum} and @{term setprod} *} haftmann@25919: haftmann@25919: text {*By Jeremy Avigad*} haftmann@25919: haftmann@25919: lemma of_nat_setsum: "of_nat (setsum f A) = (\x\A. of_nat(f x))" haftmann@25919: apply (cases "finite A") haftmann@25919: apply (erule finite_induct, auto) haftmann@25919: done haftmann@25919: haftmann@25919: lemma of_int_setsum: "of_int (setsum f A) = (\x\A. of_int(f x))" haftmann@25919: apply (cases "finite A") haftmann@25919: apply (erule finite_induct, auto) haftmann@25919: done haftmann@25919: haftmann@25919: lemma of_nat_setprod: "of_nat (setprod f A) = (\x\A. of_nat(f x))" haftmann@25919: apply (cases "finite A") haftmann@25919: apply (erule finite_induct, auto simp add: of_nat_mult) haftmann@25919: done haftmann@25919: haftmann@25919: lemma of_int_setprod: "of_int (setprod f A) = (\x\A. of_int(f x))" haftmann@25919: apply (cases "finite A") haftmann@25919: apply (erule finite_induct, auto) haftmann@25919: done haftmann@25919: haftmann@25919: lemma setprod_nonzero_nat: haftmann@25919: "finite A ==> (\x \ A. f x \ (0::nat)) ==> setprod f A \ 0" haftmann@25919: by (rule setprod_nonzero, auto) haftmann@25919: haftmann@25919: lemma setprod_zero_eq_nat: haftmann@25919: "finite A ==> (setprod f A = (0::nat)) = (\x \ A. f x = 0)" haftmann@25919: by (rule setprod_zero_eq, auto) haftmann@25919: haftmann@25919: lemma setprod_nonzero_int: haftmann@25919: "finite A ==> (\x \ A. f x \ (0::int)) ==> setprod f A \ 0" haftmann@25919: by (rule setprod_nonzero, auto) haftmann@25919: haftmann@25919: lemma setprod_zero_eq_int: haftmann@25919: "finite A ==> (setprod f A = (0::int)) = (\x \ A. f x = 0)" haftmann@25919: by (rule setprod_zero_eq, auto) haftmann@25919: haftmann@25919: lemmas int_setsum = of_nat_setsum [where 'a=int] haftmann@25919: lemmas int_setprod = of_nat_setprod [where 'a=int] haftmann@25919: haftmann@25919: haftmann@25919: subsection{*Inequality Reasoning for the Arithmetic Simproc*} haftmann@25919: haftmann@25919: lemma add_numeral_0: "Numeral0 + a = (a::'a::number_ring)" haftmann@25919: by simp haftmann@25919: haftmann@25919: lemma add_numeral_0_right: "a + Numeral0 = (a::'a::number_ring)" haftmann@25919: by simp haftmann@25919: haftmann@25919: lemma mult_numeral_1: "Numeral1 * a = (a::'a::number_ring)" haftmann@25919: by simp haftmann@25919: haftmann@25919: lemma mult_numeral_1_right: "a * Numeral1 = (a::'a::number_ring)" haftmann@25919: by simp haftmann@25919: haftmann@25919: lemma divide_numeral_1: "a / Numeral1 = (a::'a::{number_ring,field})" haftmann@25919: by simp haftmann@25919: haftmann@25919: lemma inverse_numeral_1: haftmann@25919: "inverse Numeral1 = (Numeral1::'a::{number_ring,field})" haftmann@25919: by simp haftmann@25919: haftmann@25919: text{*Theorem lists for the cancellation simprocs. The use of binary numerals haftmann@25919: for 0 and 1 reduces the number of special cases.*} haftmann@25919: haftmann@25919: lemmas add_0s = add_numeral_0 add_numeral_0_right haftmann@25919: lemmas mult_1s = mult_numeral_1 mult_numeral_1_right haftmann@25919: mult_minus1 mult_minus1_right haftmann@25919: haftmann@25919: haftmann@25919: subsection{*Special Arithmetic Rules for Abstract 0 and 1*} haftmann@25919: haftmann@25919: text{*Arithmetic computations are defined for binary literals, which leaves 0 haftmann@25919: and 1 as special cases. Addition already has rules for 0, but not 1. haftmann@25919: Multiplication and unary minus already have rules for both 0 and 1.*} haftmann@25919: haftmann@25919: haftmann@25919: lemma binop_eq: "[|f x y = g x y; x = x'; y = y'|] ==> f x' y' = g x' y'" haftmann@25919: by simp haftmann@25919: haftmann@25919: haftmann@25919: lemmas add_number_of_eq = number_of_add [symmetric] haftmann@25919: haftmann@25919: text{*Allow 1 on either or both sides*} haftmann@25919: lemma one_add_one_is_two: "1 + 1 = (2::'a::number_ring)" haftmann@25919: by (simp del: numeral_1_eq_1 add: numeral_1_eq_1 [symmetric] add_number_of_eq) haftmann@25919: haftmann@25919: lemmas add_special = haftmann@25919: one_add_one_is_two haftmann@25919: binop_eq [of "op +", OF add_number_of_eq numeral_1_eq_1 refl, standard] haftmann@25919: binop_eq [of "op +", OF add_number_of_eq refl numeral_1_eq_1, standard] haftmann@25919: haftmann@25919: text{*Allow 1 on either or both sides (1-1 already simplifies to 0)*} haftmann@25919: lemmas diff_special = haftmann@25919: binop_eq [of "op -", OF diff_number_of_eq numeral_1_eq_1 refl, standard] haftmann@25919: binop_eq [of "op -", OF diff_number_of_eq refl numeral_1_eq_1, standard] haftmann@25919: haftmann@25919: text{*Allow 0 or 1 on either side with a binary numeral on the other*} haftmann@25919: lemmas eq_special = haftmann@25919: binop_eq [of "op =", OF eq_number_of_eq numeral_0_eq_0 refl, standard] haftmann@25919: binop_eq [of "op =", OF eq_number_of_eq numeral_1_eq_1 refl, standard] haftmann@25919: binop_eq [of "op =", OF eq_number_of_eq refl numeral_0_eq_0, standard] haftmann@25919: binop_eq [of "op =", OF eq_number_of_eq refl numeral_1_eq_1, standard] haftmann@25919: haftmann@25919: text{*Allow 0 or 1 on either side with a binary numeral on the other*} haftmann@25919: lemmas less_special = haftmann@25919: binop_eq [of "op <", OF less_number_of_eq_neg numeral_0_eq_0 refl, standard] haftmann@25919: binop_eq [of "op <", OF less_number_of_eq_neg numeral_1_eq_1 refl, standard] haftmann@25919: binop_eq [of "op <", OF less_number_of_eq_neg refl numeral_0_eq_0, standard] haftmann@25919: binop_eq [of "op <", OF less_number_of_eq_neg refl numeral_1_eq_1, standard] haftmann@25919: haftmann@25919: text{*Allow 0 or 1 on either side with a binary numeral on the other*} haftmann@25919: lemmas le_special = haftmann@25919: binop_eq [of "op \", OF le_number_of_eq numeral_0_eq_0 refl, standard] haftmann@25919: binop_eq [of "op \", OF le_number_of_eq numeral_1_eq_1 refl, standard] haftmann@25919: binop_eq [of "op \", OF le_number_of_eq refl numeral_0_eq_0, standard] haftmann@25919: binop_eq [of "op \", OF le_number_of_eq refl numeral_1_eq_1, standard] haftmann@25919: haftmann@25919: lemmas arith_special[simp] = haftmann@25919: add_special diff_special eq_special less_special le_special haftmann@25919: haftmann@25919: haftmann@25919: lemma min_max_01: "min (0::int) 1 = 0 & min (1::int) 0 = 0 & haftmann@25919: max (0::int) 1 = 1 & max (1::int) 0 = 1" haftmann@25919: by(simp add:min_def max_def) haftmann@25919: haftmann@25919: lemmas min_max_special[simp] = haftmann@25919: min_max_01 haftmann@25919: max_def[of "0::int" "number_of v", standard, simp] haftmann@25919: min_def[of "0::int" "number_of v", standard, simp] haftmann@25919: max_def[of "number_of u" "0::int", standard, simp] haftmann@25919: min_def[of "number_of u" "0::int", standard, simp] haftmann@25919: max_def[of "1::int" "number_of v", standard, simp] haftmann@25919: min_def[of "1::int" "number_of v", standard, simp] haftmann@25919: max_def[of "number_of u" "1::int", standard, simp] haftmann@25919: min_def[of "number_of u" "1::int", standard, simp] haftmann@25919: haftmann@25919: text {* Legacy theorems *} haftmann@25919: haftmann@25919: lemmas zle_int = of_nat_le_iff [where 'a=int] haftmann@25919: lemmas int_int_eq = of_nat_eq_iff [where 'a=int] haftmann@25919: haftmann@25919: use "~~/src/Provers/Arith/assoc_fold.ML" haftmann@28952: use "Tools/int_arith.ML" haftmann@25919: declaration {* K int_arith_setup *} haftmann@25919: haftmann@25919: haftmann@25919: subsection{*Lemmas About Small Numerals*} haftmann@25919: haftmann@25919: lemma of_int_m1 [simp]: "of_int -1 = (-1 :: 'a :: number_ring)" haftmann@25919: proof - haftmann@25919: have "(of_int -1 :: 'a) = of_int (- 1)" by simp haftmann@25919: also have "... = - of_int 1" by (simp only: of_int_minus) haftmann@25919: also have "... = -1" by simp haftmann@25919: finally show ?thesis . haftmann@25919: qed haftmann@25919: haftmann@25919: lemma abs_minus_one [simp]: "abs (-1) = (1::'a::{ordered_idom,number_ring})" haftmann@25919: by (simp add: abs_if) haftmann@25919: haftmann@25919: lemma abs_power_minus_one [simp]: haftmann@25919: "abs(-1 ^ n) = (1::'a::{ordered_idom,number_ring,recpower})" haftmann@25919: by (simp add: power_abs) haftmann@25919: haftmann@25919: lemma of_int_number_of_eq: haftmann@25919: "of_int (number_of v) = (number_of v :: 'a :: number_ring)" haftmann@25919: by (simp add: number_of_eq) haftmann@25919: haftmann@25919: text{*Lemmas for specialist use, NOT as default simprules*} haftmann@25919: lemma mult_2: "2 * z = (z+z::'a::number_ring)" haftmann@25919: proof - haftmann@25919: have "2*z = (1 + 1)*z" by simp haftmann@25919: also have "... = z+z" by (simp add: left_distrib) haftmann@25919: finally show ?thesis . haftmann@25919: qed haftmann@25919: haftmann@25919: lemma mult_2_right: "z * 2 = (z+z::'a::number_ring)" haftmann@25919: by (subst mult_commute, rule mult_2) haftmann@25919: haftmann@25919: haftmann@25919: subsection{*More Inequality Reasoning*} haftmann@25919: haftmann@25919: lemma zless_add1_eq: "(w < z + (1::int)) = (w z) = (w z - (1::int)) = (wz)" haftmann@25919: by arith haftmann@25919: haftmann@25919: lemma int_one_le_iff_zero_less: "((1::int) \ z) = (0 < z)" haftmann@25919: by arith haftmann@25919: haftmann@25919: huffman@28958: subsection{*The functions @{term nat} and @{term int}*} haftmann@25919: haftmann@25919: text{*Simplify the terms @{term "int 0"}, @{term "int(Suc 0)"} and haftmann@25919: @{term "w + - z"}*} haftmann@25919: declare Zero_int_def [symmetric, simp] haftmann@25919: declare One_int_def [symmetric, simp] haftmann@25919: haftmann@25919: lemmas diff_int_def_symmetric = diff_int_def [symmetric, simp] haftmann@25919: haftmann@25919: lemma nat_0: "nat 0 = 0" haftmann@25919: by (simp add: nat_eq_iff) haftmann@25919: haftmann@25919: lemma nat_1: "nat 1 = Suc 0" haftmann@25919: by (subst nat_eq_iff, simp) haftmann@25919: haftmann@25919: lemma nat_2: "nat 2 = Suc (Suc 0)" haftmann@25919: by (subst nat_eq_iff, simp) haftmann@25919: haftmann@25919: lemma one_less_nat_eq [simp]: "(Suc 0 < nat z) = (1 < z)" haftmann@25919: apply (insert zless_nat_conj [of 1 z]) haftmann@25919: apply (auto simp add: nat_1) haftmann@25919: done haftmann@25919: haftmann@25919: text{*This simplifies expressions of the form @{term "int n = z"} where haftmann@25919: z is an integer literal.*} haftmann@25919: lemmas int_eq_iff_number_of [simp] = int_eq_iff [of _ "number_of v", standard] haftmann@25919: haftmann@25919: lemma split_nat [arith_split]: haftmann@25919: "P(nat(i::int)) = ((\n. i = of_nat n \ P n) & (i < 0 \ P 0))" haftmann@25919: (is "?P = (?L & ?R)") haftmann@25919: proof (cases "i < 0") haftmann@25919: case True thus ?thesis by auto haftmann@25919: next haftmann@25919: case False haftmann@25919: have "?P = ?L" haftmann@25919: proof haftmann@25919: assume ?P thus ?L using False by clarsimp haftmann@25919: next haftmann@25919: assume ?L thus ?P using False by simp haftmann@25919: qed haftmann@25919: with False show ?thesis by simp haftmann@25919: qed haftmann@25919: haftmann@25919: context ring_1 haftmann@25919: begin haftmann@25919: haftmann@25919: lemma of_int_of_nat: haftmann@25919: "of_int k = (if k < 0 then - of_nat (nat (- k)) else of_nat (nat k))" haftmann@25919: proof (cases "k < 0") haftmann@25919: case True then have "0 \ - k" by simp haftmann@25919: then have "of_nat (nat (- k)) = of_int (- k)" by (rule of_nat_nat) haftmann@25919: with True show ?thesis by simp haftmann@25919: next haftmann@25919: case False then show ?thesis by (simp add: not_less of_nat_nat) haftmann@25919: qed haftmann@25919: haftmann@25919: end haftmann@25919: haftmann@25919: lemma nat_mult_distrib: haftmann@25919: fixes z z' :: int haftmann@25919: assumes "0 \ z" haftmann@25919: shows "nat (z * z') = nat z * nat z'" haftmann@25919: proof (cases "0 \ z'") haftmann@25919: case False with assms have "z * z' \ 0" haftmann@25919: by (simp add: not_le mult_le_0_iff) haftmann@25919: then have "nat (z * z') = 0" by simp haftmann@25919: moreover from False have "nat z' = 0" by simp haftmann@25919: ultimately show ?thesis by simp haftmann@25919: next haftmann@25919: case True with assms have ge_0: "z * z' \ 0" by (simp add: zero_le_mult_iff) haftmann@25919: show ?thesis haftmann@25919: by (rule injD [of "of_nat :: nat \ int", OF inj_of_nat]) haftmann@25919: (simp only: of_nat_mult of_nat_nat [OF True] haftmann@25919: of_nat_nat [OF assms] of_nat_nat [OF ge_0], simp) haftmann@25919: qed haftmann@25919: haftmann@25919: lemma nat_mult_distrib_neg: "z \ (0::int) ==> nat(z*z') = nat(-z) * nat(-z')" haftmann@25919: apply (rule trans) haftmann@25919: apply (rule_tac [2] nat_mult_distrib, auto) haftmann@25919: done haftmann@25919: haftmann@25919: lemma nat_abs_mult_distrib: "nat (abs (w * z)) = nat (abs w) * nat (abs z)" haftmann@25919: apply (cases "z=0 | w=0") haftmann@25919: apply (auto simp add: abs_if nat_mult_distrib [symmetric] haftmann@25919: nat_mult_distrib_neg [symmetric] mult_less_0_iff) haftmann@25919: done haftmann@25919: haftmann@25919: haftmann@25919: subsection "Induction principles for int" haftmann@25919: haftmann@25919: text{*Well-founded segments of the integers*} haftmann@25919: haftmann@25919: definition haftmann@25919: int_ge_less_than :: "int => (int * int) set" haftmann@25919: where haftmann@25919: "int_ge_less_than d = {(z',z). d \ z' & z' < z}" haftmann@25919: haftmann@25919: theorem wf_int_ge_less_than: "wf (int_ge_less_than d)" haftmann@25919: proof - haftmann@25919: have "int_ge_less_than d \ measure (%z. nat (z-d))" haftmann@25919: by (auto simp add: int_ge_less_than_def) haftmann@25919: thus ?thesis haftmann@25919: by (rule wf_subset [OF wf_measure]) haftmann@25919: qed haftmann@25919: haftmann@25919: text{*This variant looks odd, but is typical of the relations suggested haftmann@25919: by RankFinder.*} haftmann@25919: haftmann@25919: definition haftmann@25919: int_ge_less_than2 :: "int => (int * int) set" haftmann@25919: where haftmann@25919: "int_ge_less_than2 d = {(z',z). d \ z & z' < z}" haftmann@25919: haftmann@25919: theorem wf_int_ge_less_than2: "wf (int_ge_less_than2 d)" haftmann@25919: proof - haftmann@25919: have "int_ge_less_than2 d \ measure (%z. nat (1+z-d))" haftmann@25919: by (auto simp add: int_ge_less_than2_def) haftmann@25919: thus ?thesis haftmann@25919: by (rule wf_subset [OF wf_measure]) haftmann@25919: qed haftmann@25919: haftmann@25919: abbreviation haftmann@25919: int :: "nat \ int" haftmann@25919: where haftmann@25919: "int \ of_nat" haftmann@25919: haftmann@25919: (* `set:int': dummy construction *) haftmann@25919: theorem int_ge_induct [case_names base step, induct set: int]: haftmann@25919: fixes i :: int haftmann@25919: assumes ge: "k \ i" and haftmann@25919: base: "P k" and haftmann@25919: step: "\i. k \ i \ P i \ P (i + 1)" haftmann@25919: shows "P i" haftmann@25919: proof - haftmann@25919: { fix n have "\i::int. n = nat(i-k) \ k \ i \ P i" haftmann@25919: proof (induct n) haftmann@25919: case 0 haftmann@25919: hence "i = k" by arith haftmann@25919: thus "P i" using base by simp haftmann@25919: next haftmann@25919: case (Suc n) haftmann@25919: then have "n = nat((i - 1) - k)" by arith haftmann@25919: moreover haftmann@25919: have ki1: "k \ i - 1" using Suc.prems by arith haftmann@25919: ultimately haftmann@25919: have "P(i - 1)" by(rule Suc.hyps) haftmann@25919: from step[OF ki1 this] show ?case by simp haftmann@25919: qed haftmann@25919: } haftmann@25919: with ge show ?thesis by fast haftmann@25919: qed haftmann@25919: haftmann@25928: (* `set:int': dummy construction *) haftmann@25928: theorem int_gr_induct [case_names base step, induct set: int]: haftmann@25919: assumes gr: "k < (i::int)" and haftmann@25919: base: "P(k+1)" and haftmann@25919: step: "\i. \k < i; P i\ \ P(i+1)" haftmann@25919: shows "P i" haftmann@25919: apply(rule int_ge_induct[of "k + 1"]) haftmann@25919: using gr apply arith haftmann@25919: apply(rule base) haftmann@25919: apply (rule step, simp+) haftmann@25919: done haftmann@25919: haftmann@25919: theorem int_le_induct[consumes 1,case_names base step]: haftmann@25919: assumes le: "i \ (k::int)" and haftmann@25919: base: "P(k)" and haftmann@25919: step: "\i. \i \ k; P i\ \ P(i - 1)" haftmann@25919: shows "P i" haftmann@25919: proof - haftmann@25919: { fix n have "\i::int. n = nat(k-i) \ i \ k \ P i" haftmann@25919: proof (induct n) haftmann@25919: case 0 haftmann@25919: hence "i = k" by arith haftmann@25919: thus "P i" using base by simp haftmann@25919: next haftmann@25919: case (Suc n) haftmann@25919: hence "n = nat(k - (i+1))" by arith haftmann@25919: moreover haftmann@25919: have ki1: "i + 1 \ k" using Suc.prems by arith haftmann@25919: ultimately haftmann@25919: have "P(i+1)" by(rule Suc.hyps) haftmann@25919: from step[OF ki1 this] show ?case by simp haftmann@25919: qed haftmann@25919: } haftmann@25919: with le show ?thesis by fast haftmann@25919: qed haftmann@25919: haftmann@25919: theorem int_less_induct [consumes 1,case_names base step]: haftmann@25919: assumes less: "(i::int) < k" and haftmann@25919: base: "P(k - 1)" and haftmann@25919: step: "\i. \i < k; P i\ \ P(i - 1)" haftmann@25919: shows "P i" haftmann@25919: apply(rule int_le_induct[of _ "k - 1"]) haftmann@25919: using less apply arith haftmann@25919: apply(rule base) haftmann@25919: apply (rule step, simp+) haftmann@25919: done haftmann@25919: haftmann@25919: subsection{*Intermediate value theorems*} haftmann@25919: haftmann@25919: lemma int_val_lemma: haftmann@25919: "(\i 1) --> haftmann@25919: f 0 \ k --> k \ f n --> (\i \ n. f i = (k::int))" haftmann@27106: apply (induct n, simp) haftmann@25919: apply (intro strip) haftmann@25919: apply (erule impE, simp) haftmann@25919: apply (erule_tac x = n in allE, simp) haftmann@25919: apply (case_tac "k = f (n+1) ") haftmann@27106: apply force haftmann@25919: apply (erule impE) haftmann@25919: apply (simp add: abs_if split add: split_if_asm) haftmann@25919: apply (blast intro: le_SucI) haftmann@25919: done haftmann@25919: haftmann@25919: lemmas nat0_intermed_int_val = int_val_lemma [rule_format (no_asm)] haftmann@25919: haftmann@25919: lemma nat_intermed_int_val: haftmann@25919: "[| \i. m \ i & i < n --> abs(f(i + 1::nat) - f i) \ 1; m < n; haftmann@25919: f m \ k; k \ f n |] ==> ? i. m \ i & i \ n & f i = (k::int)" haftmann@25919: apply (cut_tac n = "n-m" and f = "%i. f (i+m) " and k = k haftmann@25919: in int_val_lemma) haftmann@25919: apply simp haftmann@25919: apply (erule exE) haftmann@25919: apply (rule_tac x = "i+m" in exI, arith) haftmann@25919: done haftmann@25919: haftmann@25919: haftmann@25919: subsection{*Products and 1, by T. M. Rasmussen*} haftmann@25919: haftmann@25919: lemma zabs_less_one_iff [simp]: "(\z\ < 1) = (z = (0::int))" haftmann@25919: by arith haftmann@25919: haftmann@25919: lemma abs_zmult_eq_1: "(\m * n\ = 1) ==> \m\ = (1::int)" haftmann@25919: apply (cases "\n\=1") haftmann@25919: apply (simp add: abs_mult) haftmann@25919: apply (rule ccontr) haftmann@25919: apply (auto simp add: linorder_neq_iff abs_mult) haftmann@25919: apply (subgoal_tac "2 \ \m\ & 2 \ \n\") haftmann@25919: prefer 2 apply arith haftmann@25919: apply (subgoal_tac "2*2 \ \m\ * \n\", simp) haftmann@25919: apply (rule mult_mono, auto) haftmann@25919: done haftmann@25919: haftmann@25919: lemma pos_zmult_eq_1_iff_lemma: "(m * n = 1) ==> m = (1::int) | m = -1" haftmann@25919: by (insert abs_zmult_eq_1 [of m n], arith) haftmann@25919: haftmann@25919: lemma pos_zmult_eq_1_iff: "0 < (m::int) ==> (m * n = 1) = (m = 1 & n = 1)" haftmann@25919: apply (auto dest: pos_zmult_eq_1_iff_lemma) haftmann@25919: apply (simp add: mult_commute [of m]) haftmann@25919: apply (frule pos_zmult_eq_1_iff_lemma, auto) haftmann@25919: done haftmann@25919: haftmann@25919: lemma zmult_eq_1_iff: "(m*n = (1::int)) = ((m = 1 & n = 1) | (m = -1 & n = -1))" haftmann@25919: apply (rule iffI) haftmann@25919: apply (frule pos_zmult_eq_1_iff_lemma) haftmann@25919: apply (simp add: mult_commute [of m]) haftmann@25919: apply (frule pos_zmult_eq_1_iff_lemma, auto) haftmann@25919: done haftmann@25919: haftmann@25919: (* Could be simplified but Presburger only becomes available too late *) haftmann@25919: lemma infinite_UNIV_int: "~finite(UNIV::int set)" haftmann@25919: proof haftmann@25919: assume "finite(UNIV::int set)" haftmann@25919: moreover have "~(EX i::int. 2*i = 1)" haftmann@25919: by (auto simp: pos_zmult_eq_1_iff) haftmann@25919: ultimately show False using finite_UNIV_inj_surj[of "%n::int. n+n"] haftmann@25919: by (simp add:inj_on_def surj_def) (blast intro:sym) haftmann@25919: qed haftmann@25919: haftmann@25919: haftmann@25961: subsection{*Integer Powers*} haftmann@25961: haftmann@25961: instantiation int :: recpower haftmann@25961: begin haftmann@25961: haftmann@25961: primrec power_int where haftmann@25961: "p ^ 0 = (1\int)" haftmann@25961: | "p ^ (Suc n) = (p\int) * (p ^ n)" haftmann@25961: haftmann@25961: instance proof haftmann@25961: fix z :: int haftmann@25961: fix n :: nat haftmann@25961: show "z ^ 0 = 1" by simp haftmann@25961: show "z ^ Suc n = z * (z ^ n)" by simp haftmann@25961: qed haftmann@25961: haftmann@25961: end haftmann@25961: haftmann@25961: lemma zpower_zadd_distrib: "x ^ (y + z) = ((x ^ y) * (x ^ z)::int)" haftmann@25961: by (rule Power.power_add) haftmann@25961: haftmann@25961: lemma zpower_zpower: "(x ^ y) ^ z = (x ^ (y * z)::int)" haftmann@25961: by (rule Power.power_mult [symmetric]) haftmann@25961: haftmann@25961: lemma zero_less_zpower_abs_iff [simp]: haftmann@25961: "(0 < abs x ^ n) \ (x \ (0::int) | n = 0)" haftmann@25961: by (induct n) (auto simp add: zero_less_mult_iff) haftmann@25961: haftmann@25961: lemma zero_le_zpower_abs [simp]: "(0::int) \ abs x ^ n" haftmann@25961: by (induct n) (auto simp add: zero_le_mult_iff) haftmann@25961: haftmann@25961: lemma of_int_power: haftmann@25961: "of_int (z ^ n) = (of_int z ^ n :: 'a::{recpower, ring_1})" haftmann@25961: by (induct n) (simp_all add: power_Suc) haftmann@25961: haftmann@25961: lemma int_power: "int (m^n) = (int m) ^ n" haftmann@25961: by (rule of_nat_power) haftmann@25961: haftmann@25961: lemmas zpower_int = int_power [symmetric] haftmann@25961: haftmann@25919: subsection {* Configuration of the code generator *} haftmann@25919: haftmann@26507: code_datatype Pls Min Bit0 Bit1 "number_of \ int \ int" haftmann@26507: haftmann@28562: lemmas pred_succ_numeral_code [code] = haftmann@26507: pred_bin_simps succ_bin_simps haftmann@26507: haftmann@28562: lemmas plus_numeral_code [code] = haftmann@26507: add_bin_simps haftmann@26507: arith_extra_simps(1) [where 'a = int] haftmann@26507: haftmann@28562: lemmas minus_numeral_code [code] = haftmann@26507: minus_bin_simps haftmann@26507: arith_extra_simps(2) [where 'a = int] haftmann@26507: arith_extra_simps(5) [where 'a = int] haftmann@26507: haftmann@28562: lemmas times_numeral_code [code] = haftmann@26507: mult_bin_simps haftmann@26507: arith_extra_simps(4) [where 'a = int] haftmann@26507: haftmann@26507: instantiation int :: eq haftmann@26507: begin haftmann@26507: haftmann@28562: definition [code del]: "eq_class.eq k l \ k - l = (0\int)" haftmann@26507: haftmann@26507: instance by default (simp add: eq_int_def) haftmann@26507: haftmann@26507: end haftmann@26507: haftmann@28562: lemma eq_number_of_int_code [code]: haftmann@26732: "eq_class.eq (number_of k \ int) (number_of l) \ eq_class.eq k l" haftmann@26507: unfolding eq_int_def number_of_is_id .. haftmann@26507: haftmann@28562: lemma eq_int_code [code]: haftmann@26732: "eq_class.eq Int.Pls Int.Pls \ True" haftmann@26732: "eq_class.eq Int.Pls Int.Min \ False" haftmann@26732: "eq_class.eq Int.Pls (Int.Bit0 k2) \ eq_class.eq Int.Pls k2" haftmann@26732: "eq_class.eq Int.Pls (Int.Bit1 k2) \ False" haftmann@26732: "eq_class.eq Int.Min Int.Pls \ False" haftmann@26732: "eq_class.eq Int.Min Int.Min \ True" haftmann@26732: "eq_class.eq Int.Min (Int.Bit0 k2) \ False" haftmann@26732: "eq_class.eq Int.Min (Int.Bit1 k2) \ eq_class.eq Int.Min k2" huffman@28958: "eq_class.eq (Int.Bit0 k1) Int.Pls \ eq_class.eq k1 Int.Pls" haftmann@26732: "eq_class.eq (Int.Bit1 k1) Int.Pls \ False" haftmann@26732: "eq_class.eq (Int.Bit0 k1) Int.Min \ False" huffman@28958: "eq_class.eq (Int.Bit1 k1) Int.Min \ eq_class.eq k1 Int.Min" haftmann@26732: "eq_class.eq (Int.Bit0 k1) (Int.Bit0 k2) \ eq_class.eq k1 k2" haftmann@26732: "eq_class.eq (Int.Bit0 k1) (Int.Bit1 k2) \ False" haftmann@26732: "eq_class.eq (Int.Bit1 k1) (Int.Bit0 k2) \ False" haftmann@26732: "eq_class.eq (Int.Bit1 k1) (Int.Bit1 k2) \ eq_class.eq k1 k2" huffman@28958: unfolding eq_equals by simp_all haftmann@25919: haftmann@28351: lemma eq_int_refl [code nbe]: haftmann@28351: "eq_class.eq (k::int) k \ True" haftmann@28351: by (rule HOL.eq_refl) haftmann@28351: haftmann@28562: lemma less_eq_number_of_int_code [code]: haftmann@26507: "(number_of k \ int) \ number_of l \ k \ l" haftmann@26507: unfolding number_of_is_id .. haftmann@26507: haftmann@28562: lemma less_eq_int_code [code]: haftmann@26507: "Int.Pls \ Int.Pls \ True" haftmann@26507: "Int.Pls \ Int.Min \ False" haftmann@26507: "Int.Pls \ Int.Bit0 k \ Int.Pls \ k" haftmann@26507: "Int.Pls \ Int.Bit1 k \ Int.Pls \ k" haftmann@26507: "Int.Min \ Int.Pls \ True" haftmann@26507: "Int.Min \ Int.Min \ True" haftmann@26507: "Int.Min \ Int.Bit0 k \ Int.Min < k" haftmann@26507: "Int.Min \ Int.Bit1 k \ Int.Min \ k" haftmann@26507: "Int.Bit0 k \ Int.Pls \ k \ Int.Pls" haftmann@26507: "Int.Bit1 k \ Int.Pls \ k < Int.Pls" haftmann@26507: "Int.Bit0 k \ Int.Min \ k \ Int.Min" haftmann@26507: "Int.Bit1 k \ Int.Min \ k \ Int.Min" haftmann@26507: "Int.Bit0 k1 \ Int.Bit0 k2 \ k1 \ k2" haftmann@26507: "Int.Bit0 k1 \ Int.Bit1 k2 \ k1 \ k2" haftmann@26507: "Int.Bit1 k1 \ Int.Bit0 k2 \ k1 < k2" haftmann@26507: "Int.Bit1 k1 \ Int.Bit1 k2 \ k1 \ k2" huffman@28958: by simp_all haftmann@26507: haftmann@28562: lemma less_number_of_int_code [code]: haftmann@26507: "(number_of k \ int) < number_of l \ k < l" haftmann@26507: unfolding number_of_is_id .. haftmann@26507: haftmann@28562: lemma less_int_code [code]: haftmann@26507: "Int.Pls < Int.Pls \ False" haftmann@26507: "Int.Pls < Int.Min \ False" haftmann@26507: "Int.Pls < Int.Bit0 k \ Int.Pls < k" haftmann@26507: "Int.Pls < Int.Bit1 k \ Int.Pls \ k" haftmann@26507: "Int.Min < Int.Pls \ True" haftmann@26507: "Int.Min < Int.Min \ False" haftmann@26507: "Int.Min < Int.Bit0 k \ Int.Min < k" haftmann@26507: "Int.Min < Int.Bit1 k \ Int.Min < k" haftmann@26507: "Int.Bit0 k < Int.Pls \ k < Int.Pls" haftmann@26507: "Int.Bit1 k < Int.Pls \ k < Int.Pls" haftmann@26507: "Int.Bit0 k < Int.Min \ k \ Int.Min" haftmann@26507: "Int.Bit1 k < Int.Min \ k < Int.Min" haftmann@26507: "Int.Bit0 k1 < Int.Bit0 k2 \ k1 < k2" haftmann@26507: "Int.Bit0 k1 < Int.Bit1 k2 \ k1 \ k2" haftmann@26507: "Int.Bit1 k1 < Int.Bit0 k2 \ k1 < k2" haftmann@26507: "Int.Bit1 k1 < Int.Bit1 k2 \ k1 < k2" huffman@28958: by simp_all haftmann@25919: haftmann@25919: definition haftmann@25919: nat_aux :: "int \ nat \ nat" where haftmann@25919: "nat_aux i n = nat i + n" haftmann@25919: haftmann@25919: lemma [code]: haftmann@25919: "nat_aux i n = (if i \ 0 then n else nat_aux (i - 1) (Suc n))" -- {* tail recursive *} haftmann@25919: by (auto simp add: nat_aux_def nat_eq_iff linorder_not_le order_less_imp_le haftmann@25919: dest: zless_imp_add1_zle) haftmann@25919: haftmann@25919: lemma [code]: "nat i = nat_aux i 0" haftmann@25919: by (simp add: nat_aux_def) haftmann@25919: haftmann@28514: hide (open) const nat_aux haftmann@25928: haftmann@28562: lemma zero_is_num_zero [code, code inline, symmetric, code post]: haftmann@25919: "(0\int) = Numeral0" haftmann@25919: by simp haftmann@25919: haftmann@28562: lemma one_is_num_one [code, code inline, symmetric, code post]: haftmann@25919: "(1\int) = Numeral1" haftmann@25961: by simp haftmann@25919: haftmann@25919: code_modulename SML haftmann@25928: Int Integer haftmann@25919: haftmann@25919: code_modulename OCaml haftmann@25928: Int Integer haftmann@25919: haftmann@25919: code_modulename Haskell haftmann@25928: Int Integer haftmann@25919: haftmann@25919: types_code haftmann@25919: "int" ("int") haftmann@25919: attach (term_of) {* haftmann@25919: val term_of_int = HOLogic.mk_number HOLogic.intT; haftmann@25919: *} haftmann@25919: attach (test) {* haftmann@25919: fun gen_int i = haftmann@25919: let val j = one_of [~1, 1] * random_range 0 i haftmann@25919: in (j, fn () => term_of_int j) end; haftmann@25919: *} haftmann@25919: haftmann@25919: setup {* haftmann@25919: let haftmann@25919: haftmann@25919: fun strip_number_of (@{term "Int.number_of :: int => int"} $ t) = t haftmann@25919: | strip_number_of t = t; haftmann@25919: haftmann@28537: fun numeral_codegen thy defs dep module b t gr = haftmann@25919: let val i = HOLogic.dest_numeral (strip_number_of t) haftmann@25919: in haftmann@28537: SOME (Codegen.str (string_of_int i), haftmann@28537: snd (Codegen.invoke_tycodegen thy defs dep module false HOLogic.intT gr)) haftmann@25919: end handle TERM _ => NONE; haftmann@25919: haftmann@25919: in haftmann@25919: haftmann@25919: Codegen.add_codegen "numeral_codegen" numeral_codegen haftmann@25919: haftmann@25919: end haftmann@25919: *} haftmann@25919: haftmann@25919: consts_code haftmann@25919: "number_of :: int \ int" ("(_)") haftmann@25919: "0 :: int" ("0") haftmann@25919: "1 :: int" ("1") haftmann@25919: "uminus :: int => int" ("~") haftmann@25919: "op + :: int => int => int" ("(_ +/ _)") haftmann@25919: "op * :: int => int => int" ("(_ */ _)") haftmann@25919: "op \ :: int => int => bool" ("(_ <=/ _)") haftmann@25919: "op < :: int => int => bool" ("(_