# HG changeset patch # User paulson # Date 1080121829 -3600 # Node ID 0eca4aabf371a3987d46308987c7af1498845772 # Parent bdf6b7adc3ec3b81e8b546ef01756627a896848b streamlined treatment of quotients for the integers diff -r bdf6b7adc3ec -r 0eca4aabf371 src/HOL/Integ/Bin.thy --- a/src/HOL/Integ/Bin.thy Fri Mar 19 11:06:53 2004 +0100 +++ b/src/HOL/Integ/Bin.thy Wed Mar 24 10:50:29 2004 +0100 @@ -137,7 +137,7 @@ lemma odd_nonzero: "1 + z + z \ (0::int)"; proof (cases z rule: int_cases) case (nonneg n) - have le: "0 \ z+z" by (simp add: prems add_increasing) + have le: "0 \ z+z" by (simp add: nonneg add_increasing) thus ?thesis using le_imp_0_less [OF le] by (auto simp add: add_assoc) next @@ -147,7 +147,8 @@ assume eq: "1 + z + z = 0" have "0 < 1 + (int n + int n)" by (simp add: le_imp_0_less add_increasing) - also have "... = - (1 + z + z)" by (simp add: prems int_Suc add_ac) + also have "... = - (1 + z + z)" + by (simp add: neg add_assoc [symmetric], simp add: add_commute) also have "... = 0" by (simp add: eq) finally have "0<0" .. thus False by blast @@ -235,9 +236,9 @@ assume "a \ range of_int" then obtain z where a: "a = of_int z" .. hence "((1::'a) + a + a < 0) = (of_int (1 + z + z) < (of_int 0 :: 'a))" - by (simp add: prems) + by (simp add: a) also have "... = (z < 0)" by (simp only: of_int_less_iff odd_less_0) - also have "... = (a < 0)" by (simp add: prems) + also have "... = (a < 0)" by (simp add: a) finally show ?thesis . qed diff -r bdf6b7adc3ec -r 0eca4aabf371 src/HOL/Integ/IntArith.thy --- a/src/HOL/Integ/IntArith.thy Fri Mar 19 11:06:53 2004 +0100 +++ b/src/HOL/Integ/IntArith.thy Wed Mar 24 10:50:29 2004 +0100 @@ -40,17 +40,6 @@ subsection{*Inequality Reasoning for the Arithmetic Simproc*} -lemma zero_less_nat_eq [simp]: "(0 < nat z) = (0 < z)" -by (cut_tac w = 0 in zless_nat_conj, auto) - -lemma zless_imp_add1_zle: "w w + (1::int) \ z" -apply (rule eq_Abs_Integ [of z]) -apply (rule eq_Abs_Integ [of w]) -apply (simp add: linorder_not_le [symmetric] zle int_def zadd One_int_def) -done - - - lemma add_numeral_0: "Numeral0 + a = (a::'a::number_ring)" by simp @@ -170,10 +159,10 @@ lemma add1_zle_eq: "(w + (1::int) \ z) = (w z - (1::int)) = (w<(z::int))" +lemma zle_diff1_eq [simp]: "(w \ z - (1::int)) = (w(z::int))" +lemma zle_add1_eq_le [simp]: "(w < z + (1::int)) = (w\z)" by arith lemma int_one_le_iff_zero_less: "((1::int) \ z) = (0 < z)" @@ -182,32 +171,13 @@ subsection{*The Functions @{term nat} and @{term int}*} -lemma nonneg_eq_int: "[| 0 \ z; !!m. z = int m ==> P |] ==> P" -by (blast dest: nat_0_le sym) - -lemma nat_eq_iff: "(nat w = m) = (if 0 \ w then w = int m else m=0)" -by auto - -lemma nat_eq_iff2: "(m = nat w) = (if 0 \ w then w = int m else m=0)" -by auto - -lemma nat_less_iff: "0 \ w ==> (nat w < m) = (w < int m)" -apply (rule iffI) -apply (erule nat_0_le [THEN subst]) -apply (simp_all del: zless_int add: zless_int [symmetric]) -done - -lemma int_eq_iff: "(int m = z) = (m = nat z & 0 \ z)" -by (auto simp add: nat_eq_iff2) - - text{*Simplify the terms @{term "int 0"}, @{term "int(Suc 0)"} and @{term "w + - z"}*} declare Zero_int_def [symmetric, simp] declare One_int_def [symmetric, simp] text{*cooper.ML refers to this theorem*} -lemmas zdiff_def_symmetric = zdiff_def [symmetric, simp] +lemmas diff_int_def_symmetric = diff_int_def [symmetric, simp] lemma nat_0: "nat 0 = 0" by (simp add: nat_eq_iff) @@ -218,22 +188,10 @@ lemma nat_2: "nat 2 = Suc (Suc 0)" by (subst nat_eq_iff, simp) -lemma nat_less_eq_zless: "0 \ w ==> (nat w < nat z) = (w z ==> (nat w \ nat z) = (w\z)" -by (auto simp add: linorder_not_less [symmetric] zless_nat_conj) - - text{*This simplifies expressions of the form @{term "int n = z"} where z is an integer literal.*} declare int_eq_iff [of _ "number_of v", standard, simp] -lemma int_nat_eq [simp]: "int (nat z) = (if 0 \ z then z else 0)" - by simp - lemma split_nat [arith_split]: "P(nat(i::int)) = ((\n. i = int n \ P n) & (i < 0 \ P 0))" (is "?P = (?L & ?R)") @@ -251,6 +209,29 @@ qed +(*Analogous to zadd_int*) +lemma zdiff_int: "n \ m ==> int m - int n = int (m-n)" +by (induct m n rule: diff_induct, simp_all) + +lemma nat_mult_distrib: "(0::int) \ z ==> nat (z*z') = nat z * nat z'" +apply (case_tac "0 \ z'") +apply (rule inj_int [THEN injD]) +apply (simp add: zmult_int [symmetric] zero_le_mult_iff) +apply (simp add: mult_le_0_iff) +done + +lemma nat_mult_distrib_neg: "z \ (0::int) ==> nat(z*z') = nat(-z) * nat(-z')" +apply (rule trans) +apply (rule_tac [2] nat_mult_distrib, auto) +done + +lemma nat_abs_mult_distrib: "nat (abs (w * z)) = nat (abs w) * nat (abs z)" +apply (case_tac "z=0 | w=0") +apply (auto simp add: zabs_def nat_mult_distrib [symmetric] + nat_mult_distrib_neg [symmetric] mult_less_0_iff) +done + + subsubsection "Induction principles for int" (* `set:int': dummy construction *) @@ -390,42 +371,6 @@ done -subsection{*More about nat*} - -(*Analogous to zadd_int*) -lemma zdiff_int: "n \ m ==> int m - int n = int (m-n)" -by (induct m n rule: diff_induct, simp_all) - -lemma nat_add_distrib: - "[| (0::int) \ z; 0 \ z' |] ==> nat (z+z') = nat z + nat z'" -apply (rule inj_int [THEN injD]) -apply (simp add: zadd_int [symmetric]) -done - -lemma nat_diff_distrib: - "[| (0::int) \ z'; z' \ z |] ==> nat (z-z') = nat z - nat z'" -apply (rule inj_int [THEN injD]) -apply (simp add: zdiff_int [symmetric] nat_le_eq_zle) -done - -lemma nat_mult_distrib: "(0::int) \ z ==> nat (z*z') = nat z * nat z'" -apply (case_tac "0 \ z'") -apply (rule inj_int [THEN injD]) -apply (simp add: zmult_int [symmetric] zero_le_mult_iff) -apply (simp add: mult_le_0_iff) -done - -lemma nat_mult_distrib_neg: "z \ (0::int) ==> nat(z*z') = nat(-z) * nat(-z')" -apply (rule trans) -apply (rule_tac [2] nat_mult_distrib, auto) -done - -lemma nat_abs_mult_distrib: "nat (abs (w * z)) = nat (abs w) * nat (abs z)" -apply (case_tac "z=0 | w=0") -apply (auto simp add: zabs_def nat_mult_distrib [symmetric] - nat_mult_distrib_neg [symmetric] mult_less_0_iff) -done - ML {* diff -r bdf6b7adc3ec -r 0eca4aabf371 src/HOL/Integ/IntDef.thy --- a/src/HOL/Integ/IntDef.thy Fri Mar 19 11:06:53 2004 +0100 +++ b/src/HOL/Integ/IntDef.thy Wed Mar 24 10:50:29 2004 +0100 @@ -8,13 +8,15 @@ header{*The Integers as Equivalence Classes over Pairs of Natural Numbers*} theory IntDef = Equiv + NatArith: + constdefs intrel :: "((nat * nat) * (nat * nat)) set" - "intrel == {p. \x1 y1 x2 y2. p = ((x1,y1),(x2,y2)) & x1+y2 = x2+y1}" + --{*the equivalence relation underlying the integers*} + "intrel == {p. \x y u v. p = ((x,y),(u,v)) & x+v = u+y}" typedef (Integ) int = "UNIV//intrel" - by (auto simp add: quotient_def) + by (auto simp add: quotient_def) instance int :: ord .. instance int :: zero .. @@ -27,49 +29,57 @@ int :: "nat => int" "int m == Abs_Integ(intrel `` {(m,0)})" - + + defs (overloaded) - - zminus_def: "- Z == Abs_Integ(\(x,y) \ Rep_Integ(Z). intrel``{(y,x)})" Zero_int_def: "0 == int 0" One_int_def: "1 == int 1" - zadd_def: - "z + w == - Abs_Integ(\(x1,y1) \ Rep_Integ(z). \(x2,y2) \ Rep_Integ(w). - intrel``{(x1+x2, y1+y2)})" + minus_int_def: + "- z == contents (\(x,y) \ Rep_Integ(z). { Abs_Integ(intrel``{(y,x)}) })" + + add_int_def: + "z + w == + contents (\(x,y) \ Rep_Integ(z). \(u,v) \ Rep_Integ(w). + { Abs_Integ(intrel``{(x+u, y+v)}) })" + + diff_int_def: "z - (w::int) == z + (-w)" - zdiff_def: "z - (w::int) == z + (-w)" - zmult_def: - "z * w == - Abs_Integ(\(x1,y1) \ Rep_Integ(z). \(x2,y2) \ Rep_Integ(w). - intrel``{(x1*x2 + y1*y2, x1*y2 + y1*x2)})" + mult_int_def: + "z * w == + contents (\(x,y) \ Rep_Integ(z). \(u,v) \ Rep_Integ(w). + { Abs_Integ(intrel``{(x*u + y*v, x*v + y*u)}) })" + + le_int_def: + "z \ (w::int) == + \x y u v. x+v \ u+y & (x,y) \ Rep_Integ z & (u,v) \ Rep_Integ w" - zless_def: "(z < (w::int)) == (z \ w & z \ w)" + less_int_def: "(z < (w::int)) == (z \ w & z \ w)" + + +subsection{*Construction of the Integers*} - zle_def: - "z \ (w::int) == \x1 y1 x2 y2. x1 + y2 \ x2 + y1 & - (x1,y1) \ Rep_Integ z & (x2,y2) \ Rep_Integ w" +subsubsection{*Preliminary Lemmas about the Equivalence Relation*} -lemma intrel_iff [simp]: "(((x1,y1),(x2,y2)) \ intrel) = (x1+y2 = x2+y1)" -by (unfold intrel_def, blast) +lemma intrel_iff [simp]: "(((x,y),(u,v)) \ intrel) = (x+v = u+y)" +by (simp add: intrel_def) lemma equiv_intrel: "equiv UNIV intrel" -by (unfold intrel_def equiv_def refl_def sym_def trans_def, auto) +by (simp add: intrel_def equiv_def refl_def sym_def trans_def) lemmas equiv_intrel_iff = eq_equiv_class_iff [OF equiv_intrel UNIV_I UNIV_I, simp] -lemma intrel_in_integ [simp]: "intrel``{(x,y)}:Integ" -by (unfold Integ_def intrel_def quotient_def, fast) +lemma intrel_in_integ [simp]: "intrel``{(x,y)} \ Integ" +by (simp add: Integ_def intrel_def quotient_def, fast) lemma inj_on_Abs_Integ: "inj_on Abs_Integ Integ" apply (rule inj_on_inverseI) apply (erule Abs_Integ_inverse) done -declare inj_on_Abs_Integ [THEN inj_on_iff, simp] +declare inj_on_Abs_Integ [THEN inj_on_iff, simp] Abs_Integ_inverse [simp] lemma inj_Rep_Integ: "inj(Rep_Integ)" @@ -77,85 +87,64 @@ apply (rule Rep_Integ_inverse) done +text{*Case analysis on the representation of an integer as an equivalence + class.*} +lemma eq_Abs_Integ [case_names Abs_Integ, cases type: int]: + "(!!x y. z = Abs_Integ(intrel``{(x,y)}) ==> P) ==> P" +apply (rule Rep_Integ [of z, unfolded Integ_def, THEN quotientE]) +apply (drule arg_cong [where f=Abs_Integ]) +apply (auto simp add: Rep_Integ_inverse) +done -(** int: the injection from "nat" to "int" **) + +subsubsection{*@{term int}: Embedding the Naturals into the Integers*} lemma inj_int: "inj int" -apply (rule inj_onI) -apply (unfold int_def) -apply (drule inj_on_Abs_Integ [THEN inj_onD]) -apply (rule intrel_in_integ)+ -apply (drule eq_equiv_class) -apply (rule equiv_intrel, fast) -apply (simp add: intrel_def) -done +by (simp add: inj_on_def int_def) lemma int_int_eq [iff]: "(int m = int n) = (m = n)" by (fast elim!: inj_int [THEN injD]) - -subsection{*zminus: unary negation on Integ*} - -lemma zminus_congruent: "congruent intrel (%(x,y). intrel``{(y,x)})" -apply (unfold congruent_def intrel_def) -apply (auto simp add: add_ac) -done - -lemma zminus: "- Abs_Integ(intrel``{(x,y)}) = Abs_Integ(intrel `` {(y,x)})" -by (simp add: zminus_def equiv_intrel [THEN UN_equiv_class] zminus_congruent) +subsubsection{*Integer Unary Negation*} -(*Every integer can be written in the form Abs_Integ(...) *) -lemma eq_Abs_Integ: - "(!!x y. z = Abs_Integ(intrel``{(x,y)}) ==> P) ==> P" -apply (rule_tac x1=z in Rep_Integ [unfolded Integ_def, THEN quotientE]) -apply (drule_tac f = Abs_Integ in arg_cong) -apply (rule_tac p = x in PairE) -apply (simp add: Rep_Integ_inverse) -done +lemma minus: "- Abs_Integ(intrel``{(x,y)}) = Abs_Integ(intrel `` {(y,x)})" +proof - + have "congruent intrel (\(x,y). {Abs_Integ (intrel``{(y,x)})})" + by (simp add: congruent_def) + thus ?thesis + by (simp add: minus_int_def UN_equiv_class [OF equiv_intrel]) +qed -lemma zminus_zminus [simp]: "- (- z) = (z::int)" -apply (rule eq_Abs_Integ [of z]) -apply (simp add: zminus) -done +lemma zminus_zminus: "- (- z) = (z::int)" +by (cases z, simp add: minus) -lemma inj_zminus: "inj(%z::int. -z)" -apply (rule inj_onI) -apply (drule_tac f = uminus in arg_cong, simp) -done - -lemma zminus_0 [simp]: "- 0 = (0::int)" -by (simp add: int_def Zero_int_def zminus) +lemma zminus_0: "- 0 = (0::int)" +by (simp add: int_def Zero_int_def minus) -subsection{*zadd: addition on Integ*} +subsection{*Integer Addition*} -lemma zadd: - "Abs_Integ(intrel``{(x1,y1)}) + Abs_Integ(intrel``{(x2,y2)}) = - Abs_Integ(intrel``{(x1+x2, y1+y2)})" -apply (simp add: zadd_def UN_UN_split_split_eq) -apply (subst equiv_intrel [THEN UN_equiv_class2]) -apply (auto simp add: congruent2_def) -done +lemma add: + "Abs_Integ (intrel``{(x,y)}) + Abs_Integ (intrel``{(u,v)}) = + Abs_Integ (intrel``{(x+u, y+v)})" +proof - + have "congruent2 intrel + (\z w. (\(x,y). (\(u,v). {Abs_Integ (intrel `` {(x+u, y+v)})}) w) z)" + by (simp add: congruent2_def) + thus ?thesis + by (simp add: add_int_def UN_UN_split_split_eq + UN_equiv_class2 [OF equiv_intrel]) +qed -lemma zminus_zadd_distrib [simp]: "- (z + w) = (- z) + (- w::int)" -apply (rule eq_Abs_Integ [of z]) -apply (rule eq_Abs_Integ [of w]) -apply (simp add: zminus zadd) -done +lemma zminus_zadd_distrib: "- (z + w) = (- z) + (- w::int)" +by (cases z, cases w, simp add: minus add) lemma zadd_commute: "(z::int) + w = w + z" -apply (rule eq_Abs_Integ [of z]) -apply (rule eq_Abs_Integ [of w]) -apply (simp add: add_ac zadd) -done +by (cases z, cases w, simp add: add_ac add) lemma zadd_assoc: "((z1::int) + z2) + z3 = z1 + (z2 + z3)" -apply (rule eq_Abs_Integ [of z1]) -apply (rule eq_Abs_Integ [of z2]) -apply (rule eq_Abs_Integ [of z3]) -apply (simp add: zadd add_assoc) -done +by (cases z1, cases z2, cases z3, simp add: add add_assoc) (*For AC rewriting*) lemma zadd_left_commute: "x + (y + z) = y + ((x + z) ::int)" @@ -164,13 +153,12 @@ apply (rule zadd_commute) done -(*Integer addition is an AC operator*) lemmas zadd_ac = zadd_assoc zadd_commute zadd_left_commute lemmas zmult_ac = Ring_and_Field.mult_ac lemma zadd_int: "(int m) + (int n) = int (m + n)" -by (simp add: int_def zadd) +by (simp add: int_def add) lemma zadd_int_left: "(int m) + (int n + z) = int (m + n) + z" by (simp add: zadd_int zadd_assoc [symmetric]) @@ -179,104 +167,56 @@ by (simp add: One_int_def zadd_int) (*also for the instance declaration int :: plus_ac0*) -lemma zadd_0 [simp]: "(0::int) + z = z" -apply (unfold Zero_int_def int_def) -apply (rule eq_Abs_Integ [of z]) -apply (simp add: zadd) -done - -lemma zadd_0_right [simp]: "z + (0::int) = z" -by (rule trans [OF zadd_commute zadd_0]) - -lemma zadd_zminus_inverse [simp]: "z + (- z) = (0::int)" -apply (unfold int_def Zero_int_def) -apply (rule eq_Abs_Integ [of z]) -apply (simp add: zminus zadd add_commute) -done - -lemma zadd_zminus_inverse2 [simp]: "(- z) + z = (0::int)" -apply (rule zadd_commute [THEN trans]) -apply (rule zadd_zminus_inverse) +lemma zadd_0: "(0::int) + z = z" +apply (simp add: Zero_int_def int_def) +apply (cases z, simp add: add) done -lemma zadd_zminus_cancel [simp]: "z + (- z + w) = (w::int)" -by (simp add: zadd_assoc [symmetric] zadd_0) - -lemma zminus_zadd_cancel [simp]: "(-z) + (z + w) = (w::int)" -by (simp add: zadd_assoc [symmetric] zadd_0) +lemma zadd_0_right: "z + (0::int) = z" +by (rule trans [OF zadd_commute zadd_0]) -lemma zdiff0 [simp]: "(0::int) - x = -x" -by (simp add: zdiff_def) - -lemma zdiff0_right [simp]: "x - (0::int) = x" -by (simp add: zdiff_def) - -lemma zdiff_self [simp]: "x - x = (0::int)" -by (simp add: zdiff_def Zero_int_def) +lemma zadd_zminus_inverse2: "(- z) + z = (0::int)" +by (cases z, simp add: int_def Zero_int_def minus add) -(** Lemmas **) - -lemma zadd_assoc_cong: "(z::int) + v = z' + v' ==> z + (v + w) = z' + (v' + w)" -by (simp add: zadd_assoc [symmetric]) - - -subsection{*zmult: multiplication on Integ*} +subsection{*Integer Multiplication*} text{*Congruence property for multiplication*} -lemma zmult_congruent2: "congruent2 intrel - (%p1 p2. (%(x1,y1). (%(x2,y2). - intrel``{(x1*x2 + y1*y2, x1*y2 + y1*x2)}) p2) p1)" +lemma mult_congruent2: + "congruent2 intrel + (%p1 p2. (%(x,y). (%(u,v). + { Abs_Integ(intrel``{(x*u + y*v, x*v + y*u)}) }) p2) p1)" apply (rule equiv_intrel [THEN congruent2_commuteI]) - apply (force simp add: add_ac mult_ac) -apply (clarify, simp del: equiv_intrel_iff add: add_ac mult_ac) -apply (rename_tac x1 x2 y1 y2 z1 z2) -apply (rule equiv_class_eq [OF equiv_intrel intrel_iff [THEN iffD2]]) -apply (subgoal_tac "x1*z1 + y2*z1 = y1*z1 + x2*z1 & x1*z2 + y2*z2 = y1*z2 + x2*z2") -apply (simp add: mult_ac, arith) +apply (auto simp add: congruent_def mult_ac) +apply (rename_tac u v w x y z) +apply (subgoal_tac "u*y + x*y = w*y + v*y & u*z + x*z = w*z + v*z") +apply (simp add: mult_ac, arith) apply (simp add: add_mult_distrib [symmetric]) done -lemma zmult: - "Abs_Integ((intrel``{(x1,y1)})) * Abs_Integ((intrel``{(x2,y2)})) = - Abs_Integ(intrel `` {(x1*x2 + y1*y2, x1*y2 + y1*x2)})" -by (simp add: zmult_def UN_UN_split_split_eq zmult_congruent2 +lemma mult: + "Abs_Integ((intrel``{(x,y)})) * Abs_Integ((intrel``{(u,v)})) = + Abs_Integ(intrel `` {(x*u + y*v, x*v + y*u)})" +by (simp add: mult_int_def UN_UN_split_split_eq mult_congruent2 equiv_intrel [THEN UN_equiv_class2]) lemma zmult_zminus: "(- z) * w = - (z * (w::int))" -apply (rule eq_Abs_Integ [of z]) -apply (rule eq_Abs_Integ [of w]) -apply (simp add: zminus zmult add_ac) -done +by (cases z, cases w, simp add: minus mult add_ac) lemma zmult_commute: "(z::int) * w = w * z" -apply (rule eq_Abs_Integ [of z]) -apply (rule eq_Abs_Integ [of w]) -apply (simp add: zmult add_ac mult_ac) -done +by (cases z, cases w, simp add: mult add_ac mult_ac) lemma zmult_assoc: "((z1::int) * z2) * z3 = z1 * (z2 * z3)" -apply (rule eq_Abs_Integ [of z1]) -apply (rule eq_Abs_Integ [of z2]) -apply (rule eq_Abs_Integ [of z3]) -apply (simp add: add_mult_distrib2 zmult add_ac mult_ac) -done +by (cases z1, cases z2, cases z3, simp add: mult add_mult_distrib2 mult_ac) lemma zadd_zmult_distrib: "((z1::int) + z2) * w = (z1 * w) + (z2 * w)" -apply (rule eq_Abs_Integ [of z1]) -apply (rule eq_Abs_Integ [of z2]) -apply (rule eq_Abs_Integ [of w]) -apply (simp add: add_mult_distrib2 zadd zmult add_ac mult_ac) -done - -lemma zmult_zminus_right: "w * (- z) = - (w * (z::int))" -by (simp add: zmult_commute [of w] zmult_zminus) +by (cases z1, cases z2, cases w, simp add: add mult add_mult_distrib2 mult_ac) lemma zadd_zmult_distrib2: "(w::int) * (z1 + z2) = (w * z1) + (w * z2)" by (simp add: zmult_commute [of w] zadd_zmult_distrib) lemma zdiff_zmult_distrib: "((z1::int) - z2) * w = (z1 * w) - (z2 * w)" -apply (unfold zdiff_def) +apply (simp add: diff_int_def) apply (subst zadd_zmult_distrib) apply (simp add: zmult_zminus) done @@ -285,28 +225,16 @@ by (simp add: zmult_commute [of w] zdiff_zmult_distrib) lemmas int_distrib = - zadd_zmult_distrib zadd_zmult_distrib2 + zadd_zmult_distrib zadd_zmult_distrib2 zdiff_zmult_distrib zdiff_zmult_distrib2 lemma zmult_int: "(int m) * (int n) = int (m * n)" -by (simp add: int_def zmult) - -lemma zmult_0 [simp]: "0 * z = (0::int)" -apply (unfold Zero_int_def int_def) -apply (rule eq_Abs_Integ [of z]) -apply (simp add: zmult) -done +by (simp add: int_def mult) -lemma zmult_1 [simp]: "(1::int) * z = z" -apply (unfold One_int_def int_def) -apply (rule eq_Abs_Integ [of z]) -apply (simp add: zmult) -done +lemma zmult_1: "(1::int) * z = z" +by (cases z, simp add: One_int_def int_def mult) -lemma zmult_0_right [simp]: "z * 0 = (0::int)" -by (rule trans [OF zmult_commute zmult_0]) - -lemma zmult_1_right [simp]: "z * (1::int) = z" +lemma zmult_1_right: "z * (1::int) = z" by (rule trans [OF zmult_commute zmult_1]) @@ -316,46 +244,36 @@ fix i j k :: int show "(i + j) + k = i + (j + k)" by (simp add: zadd_assoc) show "i + j = j + i" by (simp add: zadd_commute) - show "0 + i = i" by simp - show "- i + i = 0" by simp - show "i - j = i + (-j)" by (simp add: zdiff_def) + show "0 + i = i" by (rule zadd_0) + show "- i + i = 0" by (rule zadd_zminus_inverse2) + show "i - j = i + (-j)" by (simp add: diff_int_def) show "(i * j) * k = i * (j * k)" by (rule zmult_assoc) show "i * j = j * i" by (rule zmult_commute) - show "1 * i = i" by simp + show "1 * i = i" by (rule zmult_1) show "(i + j) * k = i * k + j * k" by (simp add: int_distrib) - show "0 \ (1::int)" + show "0 \ (1::int)" by (simp only: Zero_int_def One_int_def One_nat_def int_int_eq) qed subsection{*The @{text "\"} Ordering*} -lemma zle: - "(Abs_Integ(intrel``{(x1,y1)}) \ Abs_Integ(intrel``{(x2,y2)})) = - (x1 + y2 \ x2 + y1)" -by (force simp add: zle_def) +lemma le: + "(Abs_Integ(intrel``{(x,y)}) \ Abs_Integ(intrel``{(u,v)})) = (x+v \ u+y)" +by (force simp add: le_int_def) lemma zle_refl: "w \ (w::int)" -apply (rule eq_Abs_Integ [of w]) -apply (force simp add: zle) -done +by (cases w, simp add: le) lemma zle_trans: "[| i \ j; j \ k |] ==> i \ (k::int)" -apply (rule eq_Abs_Integ [of i]) -apply (rule eq_Abs_Integ [of j]) -apply (rule eq_Abs_Integ [of k]) -apply (simp add: zle) -done +by (cases i, cases j, cases k, simp add: le) lemma zle_anti_sym: "[| z \ w; w \ z |] ==> z = (w::int)" -apply (rule eq_Abs_Integ [of w]) -apply (rule eq_Abs_Integ [of z]) -apply (simp add: zle) -done +by (cases w, cases z, simp add: le) (* Axiom 'order_less_le' of class 'order': *) lemma zless_le: "((w::int) < z) = (w \ z & w \ z)" -by (simp add: zless_def) +by (simp add: less_int_def) instance int :: order proof qed @@ -364,10 +282,7 @@ (* Axiom 'linorder_linear' of class 'linorder': *) lemma zle_linear: "(z::int) \ w | w \ z" -apply (rule eq_Abs_Integ [of z]) -apply (rule eq_Abs_Integ [of w]) -apply (simp add: zle linorder_linear) -done +by (cases z, cases w, simp add: le linorder_linear) instance int :: linorder proof qed (rule zle_linear) @@ -379,21 +294,8 @@ lemma int_eq_0_conv [simp]: "(int n = 0) = (n = 0)" by (simp add: Zero_int_def) -(*This lemma allows direct proofs of other <-properties*) -lemma zless_iff_Suc_zadd: - "(w < z) = (\n. z = w + int(Suc n))" -apply (rule eq_Abs_Integ [of z]) -apply (rule eq_Abs_Integ [of w]) -apply (simp add: linorder_not_le [where 'a = int, symmetric] - linorder_not_le [where 'a = nat] - zle int_def zdiff_def zadd zminus) -apply (safe dest!: less_imp_Suc_add) -apply (rule_tac x = k in exI) -apply (simp_all add: add_ac) -done - lemma zless_int [simp]: "(int m < int n) = (m j ==> k + i \ k + (j::int)" -apply (rule eq_Abs_Integ [of i]) -apply (rule eq_Abs_Integ [of j]) -apply (rule eq_Abs_Integ [of k]) -apply (simp add: zle zadd) -done +lemma zadd_left_mono: "i \ j ==> k + i \ k + (j::int)" +by (cases i, cases j, cases k, simp add: le add) -lemma zadd_strict_right_mono: "i < j ==> i + k < j + (k::int)" -apply (rule eq_Abs_Integ [of i]) -apply (rule eq_Abs_Integ [of j]) -apply (rule eq_Abs_Integ [of k]) -apply (simp add: linorder_not_le [where 'a = int, symmetric] - linorder_not_le [where 'a = nat] zle zadd) +lemma zadd_strict_right_mono: "i < j ==> i + k < j + (k::int)" +apply (cases i, cases j, cases k) +apply (simp add: linorder_not_le [where 'a = int, symmetric] + linorder_not_le [where 'a = nat] le add) done lemma zadd_zless_mono: "[| w'z |] ==> w' + z' < w + (z::int)" -by (rule order_less_le_trans [OF zadd_strict_right_mono zadd_left_mono]) +by (rule order_less_le_trans [OF zadd_strict_right_mono zadd_left_mono]) subsection{*Strict Monotonicity of Multiplication*} @@ -451,7 +348,7 @@ text{*strict, in 1st argument; proof is by induction on k>0*} lemma zmult_zless_mono2_lemma [rule_format]: "i 0 int k * i < int k * j" -apply (induct_tac "k", simp) +apply (induct_tac "k", simp) apply (simp add: int_Suc) apply (case_tac "n=0") apply (simp_all add: zadd_zmult_distrib int_Suc0_eq_1 order_le_less) @@ -459,14 +356,14 @@ done lemma zero_le_imp_eq_int: "0 \ k ==> \n. k = int n" -apply (rule eq_Abs_Integ [of k]) -apply (auto simp add: zle zadd int_def Zero_int_def) -apply (rule_tac x="x-y" in exI, simp) +apply (cases k) +apply (auto simp add: le add int_def Zero_int_def) +apply (rule_tac x="x-y" in exI, simp) done lemma zmult_zless_mono2: "[| i k*i < k*j" -apply (frule order_less_imp_le [THEN zero_le_imp_eq_int]) -apply (auto simp add: zmult_zless_mono2_lemma) +apply (frule order_less_imp_le [THEN zero_le_imp_eq_int]) +apply (auto simp add: zmult_zless_mono2_lemma) done @@ -484,49 +381,100 @@ qed +lemma zless_imp_add1_zle: "w w + (1::int) \ z" +apply (cases w, cases z) +apply (simp add: linorder_not_le [symmetric] le int_def add One_int_def) +done + subsection{*Magnitide of an Integer, as a Natural Number: @{term nat}*} constdefs nat :: "int => nat" - "nat(Z) == if Z<0 then 0 else (THE m. Z = int m)" + "nat z == contents (\(x,y) \ Rep_Integ(z). { x-y })" + +lemma nat: "nat (Abs_Integ (intrel``{(x,y)})) = x-y" +proof - + have "congruent intrel (\(x,y). {x-y})" + by (simp add: congruent_def, arith) + thus ?thesis + by (simp add: nat_def UN_equiv_class [OF equiv_intrel]) +qed lemma nat_int [simp]: "nat(int n) = n" -by (unfold nat_def, auto) +by (simp add: nat int_def) lemma nat_zero [simp]: "nat 0 = 0" -apply (unfold Zero_int_def) -apply (rule nat_int) -done +by (simp only: Zero_int_def nat_int) -lemma nat_0_le [simp]: "0 \ z ==> int (nat z) = z" -apply (rule eq_Abs_Integ [of z]) -apply (simp add: nat_def linorder_not_le [symmetric] zle int_def Zero_int_def) -apply (subgoal_tac "(THE m. x = m + y) = x-y") -apply (auto simp add: the_equality) +lemma int_nat_eq [simp]: "int (nat z) = (if 0 \ z then z else 0)" +by (cases z, simp add: nat le int_def Zero_int_def) + +corollary nat_0_le: "0 \ z ==> int (nat z) = z" +apply simp done lemma nat_le_0 [simp]: "z \ 0 ==> nat z = 0" -by (simp add: nat_def order_less_le eq_commute [of 0]) +by (cases z, simp add: nat le int_def Zero_int_def) + +lemma nat_le_eq_zle: "0 < w | 0 \ z ==> (nat w \ nat z) = (w\z)" +apply (cases w, cases z) +apply (simp add: nat le linorder_not_le [symmetric] int_def Zero_int_def, arith) +done text{*An alternative condition is @{term "0 \ w"} *} -lemma nat_mono_iff: "0 < z ==> (nat w < nat z) = (w < z)" -apply (subst zless_int [symmetric]) -apply (simp add: order_le_less) -apply (case_tac "w < 0") - apply (simp add: order_less_imp_le) -apply (simp add: linorder_not_less) +corollary nat_mono_iff: "0 < z ==> (nat w < nat z) = (w < z)" +by (simp add: nat_le_eq_zle linorder_not_le [symmetric]) + +corollary nat_less_eq_zless: "0 \ w ==> (nat w < nat z) = (w z; !!m. z = int m ==> P |] ==> P" +by (blast dest: nat_0_le sym) + +lemma nat_eq_iff: "(nat w = m) = (if 0 \ w then w = int m else m=0)" +by (cases w, simp add: nat le int_def Zero_int_def, arith) + +corollary nat_eq_iff2: "(m = nat w) = (if 0 \ w then w = int m else m=0)" +by (simp only: eq_commute [of m] nat_eq_iff) + +lemma nat_less_iff: "0 \ w ==> (nat w < m) = (w < int m)" +apply (cases w) +apply (simp add: nat le int_def Zero_int_def linorder_not_le [symmetric], arith) done +lemma int_eq_iff: "(int m = z) = (m = nat z & 0 \ z)" +by (auto simp add: nat_eq_iff2) + +lemma zero_less_nat_eq [simp]: "(0 < nat z) = (0 < z)" +by (insert zless_nat_conj [of 0], auto) + + +lemma nat_add_distrib: + "[| (0::int) \ z; 0 \ z' |] ==> nat (z+z') = nat z + nat z'" +by (cases z, cases z', simp add: nat add le int_def Zero_int_def) + +lemma nat_diff_distrib: + "[| (0::int) \ z'; z' \ z |] ==> nat (z-z') = nat z - nat z'" +by (cases z, cases z', + simp add: nat add minus diff_minus le int_def Zero_int_def) + + +lemma nat_zminus_int [simp]: "nat (- (int n)) = 0" +by (simp add: int_def minus nat Zero_int_def) + +lemma zless_nat_eq_int_zless: "(m < nat z) = (int m < z)" +by (cases z, simp add: nat le int_def linorder_not_le [symmetric], arith) + subsection{*Lemmas about the Function @{term int} and Orderings*} lemma negative_zless_0: "- (int (Suc n)) < 0" -by (simp add: zless_def) +by (simp add: order_less_le) lemma negative_zless [iff]: "- (int (Suc n)) < int m" by (rule negative_zless_0 [THEN order_less_le_trans], simp) @@ -541,10 +489,7 @@ by (subst le_minus_iff, simp) lemma int_zle_neg: "(int n \ - int m) = (n = 0 & m = 0)" -apply safe -apply (drule_tac [2] le_minus_iff [THEN iffD1]) -apply (auto dest: zle_trans [OF _ negative_zle_0]) -done +by (simp add: int_def le minus Zero_int_def) lemma not_int_zless_negative [simp]: "~ (int n < - int m)" by (simp add: linorder_not_less) @@ -553,10 +498,15 @@ by (force simp add: order_eq_iff [of "- int n"] int_zle_neg) lemma zle_iff_zadd: "(w \ z) = (\n. z = w + int n)" -by (force intro: exI [of _ "0::nat"] - intro!: not_sym [THEN not0_implies_Suc] - simp add: zless_iff_Suc_zadd order_le_less) +apply (cases w, cases z) +apply (auto simp add: le add int_def) +apply (rename_tac a b c d) +apply (rule_tac x="c+b - (a+d)" in exI) +apply arith +done +lemma abs_int_eq [simp]: "abs (int m) = int m" +by (simp add: zabs_def) text{*This version is proved for all ordered rings, not just integers! It is proved here because attribute @{text arith_split} is not available @@ -566,38 +516,6 @@ "P(abs(a::'a::ordered_ring)) = ((0 \ a --> P a) & (a < 0 --> P(-a)))" by (force dest: order_less_le_trans simp add: abs_if linorder_not_less) -lemma abs_int_eq [simp]: "abs (int m) = int m" -by (simp add: zabs_def) - - -subsection{*Misc Results*} - -lemma nat_zminus_int [simp]: "nat(- (int n)) = 0" -by (auto simp add: nat_def zero_reorient minus_less_iff) - -lemma zless_nat_eq_int_zless: "(m < nat z) = (int m < z)" -apply (case_tac "0 \ z") -apply (erule nat_0_le [THEN subst], simp) -apply (simp add: linorder_not_le) -apply (auto dest: order_less_trans simp add: order_less_imp_le) -done - -text{*A case theorem distinguishing non-negative and negative int*} - -lemma negD: "x<0 ==> \n. x = - (int (Suc n))" -by (auto simp add: zless_iff_Suc_zadd - diff_eq_eq [symmetric] zdiff_def) - -lemma int_cases [cases type: int, case_names nonneg neg]: - "[|!! n. z = int n ==> P; !! n. z = - (int (Suc n)) ==> P |] ==> P" -apply (case_tac "z < 0", blast dest!: negD) -apply (simp add: linorder_not_less) -apply (blast dest: nat_0_le [THEN sym]) -done - -lemma int_induct [induct type: int, case_names nonneg neg]: - "[|!! n. P (int n); !!n. P (- (int (Suc n))) |] ==> P z" - by (cases z) auto subsection{*The Constants @{term neg} and @{term iszero}*} @@ -610,7 +528,7 @@ (*For simplifying equalities*) iszero :: "'a::semiring => bool" "iszero z == z = (0)" - + lemma not_neg_int [simp]: "~ neg(int n)" by (simp add: neg_def) @@ -623,22 +541,23 @@ lemma not_neg_eq_ge_0: "(~neg x) = (0 \ x)" by (simp add: neg_def linorder_not_less) + subsection{*To simplify inequalities when Numeral1 can get simplified to 1*} lemma not_neg_0: "~ neg 0" by (simp add: One_int_def neg_def) lemma not_neg_1: "~ neg 1" -by (simp add: neg_def linorder_not_less zero_le_one) +by (simp add: neg_def linorder_not_less zero_le_one) lemma iszero_0: "iszero 0" by (simp add: iszero_def) lemma not_iszero_1: "~ iszero 1" -by (simp add: iszero_def eq_commute) +by (simp add: iszero_def eq_commute) lemma neg_nat: "neg z ==> nat z = 0" -by (simp add: nat_def neg_def) +by (simp add: neg_def order_less_imp_le) lemma not_neg_nat: "~ neg z ==> int (nat z) = z" by (simp add: linorder_not_less neg_def) @@ -657,36 +576,36 @@ lemma of_nat_add [simp]: "of_nat (m+n) = of_nat m + of_nat n" apply (induct m) -apply (simp_all add: add_ac) +apply (simp_all add: add_ac) done lemma of_nat_mult [simp]: "of_nat (m*n) = of_nat m * of_nat n" -apply (induct m) -apply (simp_all add: mult_ac add_ac right_distrib) +apply (induct m) +apply (simp_all add: mult_ac add_ac right_distrib) done lemma zero_le_imp_of_nat: "0 \ (of_nat m::'a::ordered_semiring)" -apply (induct m, simp_all) -apply (erule order_trans) -apply (rule less_add_one [THEN order_less_imp_le]) +apply (induct m, simp_all) +apply (erule order_trans) +apply (rule less_add_one [THEN order_less_imp_le]) done lemma less_imp_of_nat_less: "m < n ==> of_nat m < (of_nat n::'a::ordered_semiring)" -apply (induct m n rule: diff_induct, simp_all) -apply (insert add_le_less_mono [OF zero_le_imp_of_nat zero_less_one], force) +apply (induct m n rule: diff_induct, simp_all) +apply (insert add_le_less_mono [OF zero_le_imp_of_nat zero_less_one], force) done lemma of_nat_less_imp_less: "of_nat m < (of_nat n::'a::ordered_semiring) ==> m < n" -apply (induct m n rule: diff_induct, simp_all) -apply (insert zero_le_imp_of_nat) -apply (force simp add: linorder_not_less [symmetric]) +apply (induct m n rule: diff_induct, simp_all) +apply (insert zero_le_imp_of_nat) +apply (force simp add: linorder_not_less [symmetric]) done lemma of_nat_less_iff [simp]: "(of_nat m < (of_nat n::'a::ordered_semiring)) = (m (of_nat n::'a::ordered_semiring)) = (m \ n)" -by (simp add: linorder_not_less [symmetric]) +by (simp add: linorder_not_less [symmetric]) text{*Special cases where either operand is zero*} declare of_nat_le_iff [of 0, simplified, simp] @@ -704,7 +623,7 @@ a finite field, which indeed wraps back to zero.*} lemma of_nat_eq_iff [simp]: "(of_nat m = (of_nat n::'a::ordered_semiring)) = (m = n)" -by (simp add: order_eq_iff) +by (simp add: order_eq_iff) text{*Special cases where either operand is zero*} declare of_nat_eq_iff [of 0, simplified, simp] @@ -713,7 +632,7 @@ lemma of_nat_diff [simp]: "n \ m ==> of_nat (m - n) = of_nat m - (of_nat n :: 'a::ring)" by (simp del: of_nat_add - add: compare_rls of_nat_add [symmetric] split add: nat_diff_split) + add: compare_rls of_nat_add [symmetric] split add: nat_diff_split) subsection{*The Set of Natural Numbers*} @@ -725,29 +644,29 @@ syntax (xsymbols) Nats :: "'a set" ("\") lemma of_nat_in_Nats [simp]: "of_nat n \ Nats" -by (simp add: Nats_def) +by (simp add: Nats_def) lemma Nats_0 [simp]: "0 \ Nats" -apply (simp add: Nats_def) -apply (rule range_eqI) +apply (simp add: Nats_def) +apply (rule range_eqI) apply (rule of_nat_0 [symmetric]) done lemma Nats_1 [simp]: "1 \ Nats" -apply (simp add: Nats_def) -apply (rule range_eqI) +apply (simp add: Nats_def) +apply (rule range_eqI) apply (rule of_nat_1 [symmetric]) done lemma Nats_add [simp]: "[|a \ Nats; b \ Nats|] ==> a+b \ Nats" -apply (auto simp add: Nats_def) -apply (rule range_eqI) +apply (auto simp add: Nats_def) +apply (rule range_eqI) apply (rule of_nat_add [symmetric]) done lemma Nats_mult [simp]: "[|a \ Nats; b \ Nats|] ==> a*b \ Nats" -apply (auto simp add: Nats_def) -apply (rule range_eqI) +apply (auto simp add: Nats_def) +apply (rule range_eqI) apply (rule of_nat_mult [symmetric]) done @@ -755,7 +674,7 @@ lemma int_eq_of_nat: "int = (of_nat :: nat => int)" proof fix n - show "int n = of_nat n" by (induct n, simp_all add: int_Suc add_ac) + show "int n = of_nat n" by (induct n, simp_all add: int_Suc add_ac) qed @@ -769,9 +688,9 @@ lemma of_int: "of_int (Abs_Integ (intrel `` {(i,j)})) = of_nat i - of_nat j" apply (simp add: of_int_def) -apply (rule the_equality, auto) +apply (rule the_equality, auto) apply (simp add: compare_rls add_ac of_nat_add [symmetric] - del: of_nat_add) + del: of_nat_add) done lemma of_int_0 [simp]: "of_int 0 = 0" @@ -781,32 +700,26 @@ by (simp add: of_int One_int_def int_def) lemma of_int_add [simp]: "of_int (w+z) = of_int w + of_int z" -apply (rule eq_Abs_Integ [of w]) -apply (rule eq_Abs_Integ [of z]) -apply (simp add: compare_rls of_int zadd) -done +by (cases w, cases z, simp add: compare_rls of_int add) lemma of_int_minus [simp]: "of_int (-z) = - (of_int z)" -apply (rule eq_Abs_Integ [of z]) -apply (simp add: compare_rls of_int zminus) -done +by (cases z, simp add: compare_rls of_int minus) lemma of_int_diff [simp]: "of_int (w-z) = of_int w - of_int z" by (simp add: diff_minus) lemma of_int_mult [simp]: "of_int (w*z) = of_int w * of_int z" -apply (rule eq_Abs_Integ [of w]) -apply (rule eq_Abs_Integ [of z]) -apply (simp add: compare_rls of_int left_diff_distrib right_diff_distrib - zmult add_ac) +apply (cases w, cases z) +apply (simp add: compare_rls of_int left_diff_distrib right_diff_distrib + mult add_ac) done lemma of_int_le_iff [simp]: "(of_int w \ (of_int z::'a::ordered_ring)) = (w \ z)" -apply (rule eq_Abs_Integ [of w]) -apply (rule eq_Abs_Integ [of z]) -apply (simp add: compare_rls of_int zle zdiff_def zadd zminus - of_nat_add [symmetric] del: of_nat_add) +apply (cases w) +apply (cases z) +apply (simp add: compare_rls of_int le diff_int_def add minus + of_nat_add [symmetric] del: of_nat_add) done text{*Special cases where either operand is zero*} @@ -824,7 +737,7 @@ text{*The ordering on the ring is necessary. See @{text of_nat_eq_iff} above.*} lemma of_int_eq_iff [simp]: "(of_int w = (of_int z::'a::ordered_ring)) = (w = z)" -by (simp add: order_eq_iff) +by (simp add: order_eq_iff) text{*Special cases where either operand is zero*} declare of_int_eq_iff [of 0, simplified, simp] @@ -842,52 +755,52 @@ Ints :: "'a set" ("\") lemma Ints_0 [simp]: "0 \ Ints" -apply (simp add: Ints_def) -apply (rule range_eqI) +apply (simp add: Ints_def) +apply (rule range_eqI) apply (rule of_int_0 [symmetric]) done lemma Ints_1 [simp]: "1 \ Ints" -apply (simp add: Ints_def) -apply (rule range_eqI) +apply (simp add: Ints_def) +apply (rule range_eqI) apply (rule of_int_1 [symmetric]) done lemma Ints_add [simp]: "[|a \ Ints; b \ Ints|] ==> a+b \ Ints" -apply (auto simp add: Ints_def) -apply (rule range_eqI) +apply (auto simp add: Ints_def) +apply (rule range_eqI) apply (rule of_int_add [symmetric]) done lemma Ints_minus [simp]: "a \ Ints ==> -a \ Ints" -apply (auto simp add: Ints_def) -apply (rule range_eqI) +apply (auto simp add: Ints_def) +apply (rule range_eqI) apply (rule of_int_minus [symmetric]) done lemma Ints_diff [simp]: "[|a \ Ints; b \ Ints|] ==> a-b \ Ints" -apply (auto simp add: Ints_def) -apply (rule range_eqI) +apply (auto simp add: Ints_def) +apply (rule range_eqI) apply (rule of_int_diff [symmetric]) done lemma Ints_mult [simp]: "[|a \ Ints; b \ Ints|] ==> a*b \ Ints" -apply (auto simp add: Ints_def) -apply (rule range_eqI) +apply (auto simp add: Ints_def) +apply (rule range_eqI) apply (rule of_int_mult [symmetric]) done text{*Collapse nested embeddings*} lemma of_int_of_nat_eq [simp]: "of_int (of_nat n) = of_nat n" -by (induct n, auto) +by (induct n, auto) lemma of_int_int_eq [simp]: "of_int (int n) = int n" -by (simp add: int_eq_of_nat) +by (simp add: int_eq_of_nat) lemma Ints_cases [case_names of_int, cases set: Ints]: "q \ \ ==> (!!z. q = of_int z ==> C) ==> C" -proof (unfold Ints_def) +proof (simp add: Ints_def) assume "!!z. q = of_int z ==> C" assume "q \ range of_int" thus C .. qed @@ -963,11 +876,41 @@ by (rule setprod_zero_eq, auto) +text{*Now we replace the case analysis rule by a more conventional one: +whether an integer is negative or not.*} + +lemma zless_iff_Suc_zadd: + "(w < z) = (\n. z = w + int(Suc n))" +apply (cases z, cases w) +apply (auto simp add: le add int_def linorder_not_le [symmetric]) +apply (rename_tac a b c d) +apply (rule_tac x="a+d - Suc(c+b)" in exI) +apply arith +done + +lemma negD: "x<0 ==> \n. x = - (int (Suc n))" +apply (cases x) +apply (auto simp add: le minus Zero_int_def int_def order_less_le) +apply (rule_tac x="y - Suc x" in exI) +apply arith +done + +theorem int_cases [cases type: int, case_names nonneg neg]: + "[|!! n. z = int n ==> P; !! n. z = - (int (Suc n)) ==> P |] ==> P" +apply (case_tac "z < 0", blast dest!: negD) +apply (simp add: linorder_not_less) +apply (blast dest: nat_0_le [THEN sym]) +done + +theorem int_induct [induct type: int, case_names nonneg neg]: + "[|!! n. P (int n); !!n. P (- (int (Suc n))) |] ==> P z" + by (cases z) auto + + (*Legacy ML bindings, but no longer the structure Int.*) ML {* val zabs_def = thm "zabs_def" -val nat_def = thm "nat_def" val int_0 = thm "int_0"; val int_1 = thm "int_1"; @@ -1002,11 +945,7 @@ val int_def = thm "int_def"; val Zero_int_def = thm "Zero_int_def"; val One_int_def = thm "One_int_def"; -val zadd_def = thm "zadd_def"; -val zdiff_def = thm "zdiff_def"; -val zless_def = thm "zless_def"; -val zle_def = thm "zle_def"; -val zmult_def = thm "zmult_def"; +val diff_int_def = thm "diff_int_def"; val intrel_iff = thm "intrel_iff"; val equiv_intrel = thm "equiv_intrel"; @@ -1015,13 +954,8 @@ val inj_on_Abs_Integ = thm "inj_on_Abs_Integ"; val inj_Rep_Integ = thm "inj_Rep_Integ"; val inj_int = thm "inj_int"; -val zminus_congruent = thm "zminus_congruent"; -val zminus = thm "zminus"; -val eq_Abs_Integ = thm "eq_Abs_Integ"; val zminus_zminus = thm "zminus_zminus"; -val inj_zminus = thm "inj_zminus"; val zminus_0 = thm "zminus_0"; -val zadd = thm "zadd"; val zminus_zadd_distrib = thm "zminus_zadd_distrib"; val zadd_commute = thm "zadd_commute"; val zadd_assoc = thm "zadd_assoc"; @@ -1033,30 +967,17 @@ val int_Suc = thm "int_Suc"; val zadd_0 = thm "zadd_0"; val zadd_0_right = thm "zadd_0_right"; -val zadd_zminus_inverse = thm "zadd_zminus_inverse"; -val zadd_zminus_inverse2 = thm "zadd_zminus_inverse2"; -val zadd_zminus_cancel = thm "zadd_zminus_cancel"; -val zminus_zadd_cancel = thm "zminus_zadd_cancel"; -val zdiff0 = thm "zdiff0"; -val zdiff0_right = thm "zdiff0_right"; -val zdiff_self = thm "zdiff_self"; -val zmult_congruent2 = thm "zmult_congruent2"; -val zmult = thm "zmult"; val zmult_zminus = thm "zmult_zminus"; val zmult_commute = thm "zmult_commute"; val zmult_assoc = thm "zmult_assoc"; val zadd_zmult_distrib = thm "zadd_zmult_distrib"; -val zmult_zminus_right = thm "zmult_zminus_right"; val zadd_zmult_distrib2 = thm "zadd_zmult_distrib2"; val zdiff_zmult_distrib = thm "zdiff_zmult_distrib"; val zdiff_zmult_distrib2 = thm "zdiff_zmult_distrib2"; val int_distrib = thms "int_distrib"; val zmult_int = thm "zmult_int"; -val zmult_0 = thm "zmult_0"; val zmult_1 = thm "zmult_1"; -val zmult_0_right = thm "zmult_0_right"; val zmult_1_right = thm "zmult_1_right"; -val zless_iff_Suc_zadd = thm "zless_iff_Suc_zadd"; val int_int_eq = thm "int_int_eq"; val int_eq_0_conv = thm "int_eq_0_conv"; val zless_int = thm "zless_int"; diff -r bdf6b7adc3ec -r 0eca4aabf371 src/HOL/Integ/IntDiv.thy --- a/src/HOL/Integ/IntDiv.thy Fri Mar 19 11:06:53 2004 +0100 +++ b/src/HOL/Integ/IntDiv.thy Wed Mar 24 10:50:29 2004 +0100 @@ -95,12 +95,12 @@ "[| b*q' + r' \ b*q + r; 0 \ r'; 0 < b; r < b |] ==> q' \ (q::int)" apply (subgoal_tac "r' + b * (q'-q) \ r") - prefer 2 apply (simp add: zdiff_zmult_distrib2) + prefer 2 apply (simp add: right_diff_distrib) apply (subgoal_tac "0 < b * (1 + q - q') ") apply (erule_tac [2] order_le_less_trans) - prefer 2 apply (simp add: zdiff_zmult_distrib2 zadd_zmult_distrib2) + prefer 2 apply (simp add: right_diff_distrib right_distrib) apply (subgoal_tac "b * q' < b * (1 + q) ") - prefer 2 apply (simp add: zdiff_zmult_distrib2 zadd_zmult_distrib2) + prefer 2 apply (simp add: right_diff_distrib right_distrib) apply (simp add: mult_less_cancel_left) done @@ -158,7 +158,7 @@ (*main argument*) apply (subst posDivAlg_eqn, simp_all) apply (erule splitE) -apply (auto simp add: zadd_zmult_distrib2 Let_def) +apply (auto simp add: right_distrib Let_def) done @@ -186,7 +186,7 @@ (*main argument*) apply (subst negDivAlg_eqn, assumption) apply (erule splitE) -apply (auto simp add: zadd_zmult_distrib2 Let_def) +apply (auto simp add: right_distrib Let_def) done @@ -321,7 +321,7 @@ "quorem((a,b),(q,r)) ==> quorem ((-a,b), (if r=0 then -q else -q - 1), (if r=0 then 0 else b-r))" -by (force simp add: split_ifs quorem_def linorder_neq_iff zdiff_zmult_distrib2) +by (force simp add: split_ifs quorem_def linorder_neq_iff right_diff_distrib) lemma zdiv_zminus1_eq_if: @@ -363,7 +363,7 @@ lemma self_quotient_aux2: "[| (0::int) < a; a = r + a*q; 0 \ r |] ==> q \ 1" apply (subgoal_tac "0 \ a* (1-q) ") apply (simp add: zero_le_mult_iff) -apply (simp add: zdiff_zmult_distrib2) +apply (simp add: right_diff_distrib) done lemma self_quotient: "[| quorem((a,a),(q,r)); a ~= (0::int) |] ==> q = 1" @@ -517,7 +517,7 @@ "[| 0 \ b'*q' + r'; r' < b'; 0 < b' |] ==> 0 \ (q'::int)" apply (subgoal_tac "0 < b'* (q' + 1) ") apply (simp add: zero_less_mult_iff) -apply (simp add: zadd_zmult_distrib2) +apply (simp add: right_distrib) done lemma zdiv_mono2_lemma: @@ -529,7 +529,7 @@ apply (simp add: mult_less_cancel_left) apply (subgoal_tac "b*q = r' - r + b'*q'") prefer 2 apply simp -apply (simp (no_asm_simp) add: zadd_zmult_distrib2) +apply (simp (no_asm_simp) add: right_distrib) apply (subst zadd_commute, rule zadd_zless_mono, arith) apply (rule mult_right_mono, auto) done @@ -559,7 +559,7 @@ apply (frule q_neg_lemma, assumption+) apply (subgoal_tac "b*q' < b* (q + 1) ") apply (simp add: mult_less_cancel_left) -apply (simp add: zadd_zmult_distrib2) +apply (simp add: right_distrib) apply (subgoal_tac "b*q' \ b'*q'") prefer 2 apply (simp add: mult_right_mono_neg) apply (subgoal_tac "b'*q' < b + b*q") @@ -585,7 +585,7 @@ lemma zmult1_lemma: "[| quorem((b,c),(q,r)); c ~= 0 |] ==> quorem ((a*b, c), (a*q + a*r div c, a*r mod c))" -by (force simp add: split_ifs quorem_def linorder_neq_iff zadd_zmult_distrib2) +by (force simp add: split_ifs quorem_def linorder_neq_iff right_distrib) lemma zdiv_zmult1_eq: "(a*b) div c = a*(b div c) + a*(b mod c) div (c::int)" apply (case_tac "c = 0", simp add: DIVISION_BY_ZERO) @@ -637,7 +637,7 @@ lemma zadd1_lemma: "[| quorem((a,c),(aq,ar)); quorem((b,c),(bq,br)); c ~= 0 |] ==> quorem ((a+b, c), (aq + bq + (ar+br) div c, (ar+br) mod c))" -by (force simp add: split_ifs quorem_def linorder_neq_iff zadd_zmult_distrib2) +by (force simp add: split_ifs quorem_def linorder_neq_iff right_distrib) (*NOT suitable for rewriting: the RHS has an instance of the LHS*) lemma zdiv_zadd1_eq: @@ -700,7 +700,7 @@ lemma zmult2_lemma_aux1: "[| (0::int) < c; b < r; r \ 0 |] ==> b*c < b*(q mod c) + r" apply (subgoal_tac "b * (c - q mod c) < r * 1") -apply (simp add: zdiff_zmult_distrib2) +apply (simp add: right_diff_distrib) apply (rule order_le_less_trans) apply (erule_tac [2] mult_strict_right_mono) apply (rule mult_left_mono_neg) @@ -722,7 +722,7 @@ lemma zmult2_lemma_aux4: "[| (0::int) < c; 0 \ r; r < b |] ==> b * (q mod c) + r < b * c" apply (subgoal_tac "r * 1 < b * (c - q mod c) ") -apply (simp add: zdiff_zmult_distrib2) +apply (simp add: right_diff_distrib) apply (rule order_less_le_trans) apply (erule mult_strict_right_mono) apply (rule_tac [2] mult_left_mono) @@ -733,7 +733,7 @@ lemma zmult2_lemma: "[| quorem ((a,b), (q,r)); b ~= 0; 0 < c |] ==> quorem ((a, b*c), (q div c, b*(q mod c) + r))" by (auto simp add: mult_ac quorem_def linorder_neq_iff - zero_less_mult_iff zadd_zmult_distrib2 [symmetric] + zero_less_mult_iff right_distrib [symmetric] zmult2_lemma_aux1 zmult2_lemma_aux2 zmult2_lemma_aux3 zmult2_lemma_aux4) lemma zdiv_zmult2_eq: "(0::int) < c ==> a div (b*c) = (a div b) div c" @@ -887,9 +887,9 @@ lemma neg_zdiv_mult_2: "a \ (0::int) ==> (1 + 2*b) div (2*a) = (b+1) div a" apply (subgoal_tac " (1 + 2* (-b - 1)) div (2 * (-a)) = (-b - 1) div (-a) ") apply (rule_tac [2] pos_zdiv_mult_2) -apply (auto simp add: zmult_zminus_right right_diff_distrib) +apply (auto simp add: minus_mult_right [symmetric] right_diff_distrib) apply (subgoal_tac " (-1 - (2 * b)) = - (1 + (2 * b))") -apply (simp only: zdiv_zminus_zminus zdiff_def zminus_zadd_distrib [symmetric], +apply (simp only: zdiv_zminus_zminus diff_minus minus_add_distrib [symmetric], simp) done @@ -934,16 +934,15 @@ apply (simp) done - lemma neg_zmod_mult_2: "a \ (0::int) ==> (1 + 2*b) mod (2*a) = 2 * ((b+1) mod a) - 1" apply (subgoal_tac "(1 + 2* (-b - 1)) mod (2* (-a)) = 1 + 2* ((-b - 1) mod (-a))") apply (rule_tac [2] pos_zmod_mult_2) -apply (auto simp add: zmult_zminus_right right_diff_distrib) +apply (auto simp add: minus_mult_right [symmetric] right_diff_distrib) apply (subgoal_tac " (-1 - (2 * b)) = - (1 + (2 * b))") prefer 2 apply simp -apply (simp only: zmod_zminus_zminus zdiff_def zminus_zadd_distrib [symmetric]) +apply (simp only: zmod_zminus_zminus diff_minus minus_add_distrib [symmetric]) done lemma zmod_number_of_BIT [simp]: @@ -1003,7 +1002,7 @@ lemma zdvd_0_right [iff]: "(m::int) dvd 0" apply (unfold dvd_def) - apply (blast intro: zmult_0_right [symmetric]) + apply (blast intro: mult_zero_right [symmetric]) done lemma zdvd_0_left [iff]: "(0 dvd (m::int)) = (m = 0)" @@ -1040,12 +1039,12 @@ lemma zdvd_zadd: "k dvd m ==> k dvd n ==> k dvd (m + n :: int)" apply (unfold dvd_def) - apply (blast intro: zadd_zmult_distrib2 [symmetric]) + apply (blast intro: right_distrib [symmetric]) done lemma zdvd_zdiff: "k dvd m ==> k dvd n ==> k dvd (m - n :: int)" apply (unfold dvd_def) - apply (blast intro: zdiff_zmult_distrib2 [symmetric]) + apply (blast intro: right_diff_distrib [symmetric]) done lemma zdvd_zdiffD: "k dvd m - n ==> k dvd n ==> k dvd (m::int)" diff -r bdf6b7adc3ec -r 0eca4aabf371 src/HOL/Integ/IntDiv_setup.ML --- a/src/HOL/Integ/IntDiv_setup.ML Fri Mar 19 11:06:53 2004 +0100 +++ b/src/HOL/Integ/IntDiv_setup.ML Wed Mar 24 10:50:29 2004 +0100 @@ -24,7 +24,7 @@ val trans = trans val prove_eq_sums = - let val simps = zdiff_def :: Int_Numeral_Simprocs.add_0s @ zadd_ac + let val simps = diff_int_def :: Int_Numeral_Simprocs.add_0s @ zadd_ac in NatArithUtils.prove_conv all_tac (NatArithUtils.simp_all simps) end end; diff -r bdf6b7adc3ec -r 0eca4aabf371 src/HOL/Integ/cooper_proof.ML --- a/src/HOL/Integ/cooper_proof.ML Fri Mar 19 11:06:53 2004 +0100 +++ b/src/HOL/Integ/cooper_proof.ML Wed Mar 24 10:50:29 2004 +0100 @@ -44,7 +44,7 @@ (*-----------------------------------------------------------------*) val presburger_ss = simpset_of (theory "Presburger") - addsimps [zdiff_def] delsimps [thm"zdiff_def_symmetric"]; + addsimps [diff_int_def] delsimps [thm"diff_int_def_symmetric"]; val cboolT = ctyp_of (sign_of HOL.thy) HOLogic.boolT; (*Theorems that will be used later for the proofgeneration*) diff -r bdf6b7adc3ec -r 0eca4aabf371 src/HOL/Set.thy --- a/src/HOL/Set.thy Fri Mar 19 11:06:53 2004 +0100 +++ b/src/HOL/Set.thy Wed Mar 24 10:50:29 2004 +0100 @@ -1949,6 +1949,16 @@ by blast +subsection {* Getting the Contents of a Singleton Set *} + +constdefs + contents :: "'a set => 'a" + "contents X == THE x. X = {x}" + +lemma contents_eq [simp]: "contents {x} = x" +by (simp add: contents_def) + + subsection {* Transitivity rules for calculational reasoning *} lemma forw_subst: "a = b ==> P b ==> P a" diff -r bdf6b7adc3ec -r 0eca4aabf371 src/HOL/Tools/Presburger/cooper_proof.ML --- a/src/HOL/Tools/Presburger/cooper_proof.ML Fri Mar 19 11:06:53 2004 +0100 +++ b/src/HOL/Tools/Presburger/cooper_proof.ML Wed Mar 24 10:50:29 2004 +0100 @@ -44,7 +44,7 @@ (*-----------------------------------------------------------------*) val presburger_ss = simpset_of (theory "Presburger") - addsimps [zdiff_def] delsimps [thm"zdiff_def_symmetric"]; + addsimps [diff_int_def] delsimps [thm"diff_int_def_symmetric"]; val cboolT = ctyp_of (sign_of HOL.thy) HOLogic.boolT; (*Theorems that will be used later for the proofgeneration*)