src/HOL/Integ/Integ.ML
changeset 925 15539deb6863
child 972 e61b058d58d2
equal deleted inserted replaced
924:806721cfbf46 925:15539deb6863
       
     1 (*  Title: 	Integ.ML
       
     2     ID:         $Id$
       
     3     Authors: 	Riccardo Mattolini, Dip. Sistemi e Informatica
       
     4         	Lawrence C Paulson, Cambridge University Computer Laboratory
       
     5     Copyright   1994 Universita' di Firenze
       
     6     Copyright   1993  University of Cambridge
       
     7 
       
     8 The integers as equivalence classes over nat*nat.
       
     9 
       
    10 Could also prove...
       
    11 "znegative(z) ==> $# zmagnitude(z) = $~ z"
       
    12 "~ znegative(z) ==> $# zmagnitude(z) = z"
       
    13 < is a linear ordering
       
    14 + and * are monotonic wrt <
       
    15 *)
       
    16 
       
    17 open Integ;
       
    18 
       
    19 
       
    20 (*** Proving that intrel is an equivalence relation ***)
       
    21 
       
    22 val eqa::eqb::prems = goal Arith.thy 
       
    23     "[| (x1::nat) + y2 = x2 + y1; x2 + y3 = x3 + y2 |] ==> \
       
    24 \       x1 + y3 = x3 + y1";
       
    25 by (res_inst_tac [("k2","x2")] (add_left_cancel RS iffD1) 1);
       
    26 by (rtac (add_left_commute RS trans) 1);
       
    27 by (rtac (eqb RS ssubst) 1);
       
    28 by (rtac (add_left_commute RS trans) 1);
       
    29 by (rtac (eqa RS ssubst) 1);
       
    30 by (rtac (add_left_commute) 1);
       
    31 qed "integ_trans_lemma";
       
    32 
       
    33 (** Natural deduction for intrel **)
       
    34 
       
    35 val prems = goalw Integ.thy [intrel_def]
       
    36     "[| x1+y2 = x2+y1|] ==> \
       
    37 \    <<x1,y1>,<x2,y2>>: intrel";
       
    38 by (fast_tac (rel_cs addIs prems) 1);
       
    39 qed "intrelI";
       
    40 
       
    41 (*intrelE is hard to derive because fast_tac tries hyp_subst_tac so soon*)
       
    42 goalw Integ.thy [intrel_def]
       
    43   "p: intrel --> (EX x1 y1 x2 y2. \
       
    44 \                  p = <<x1,y1>,<x2,y2>> & x1+y2 = x2+y1)";
       
    45 by (fast_tac rel_cs 1);
       
    46 qed "intrelE_lemma";
       
    47 
       
    48 val [major,minor] = goal Integ.thy
       
    49   "[| p: intrel;  \
       
    50 \     !!x1 y1 x2 y2. [| p = <<x1,y1>,<x2,y2>>;  x1+y2 = x2+y1|] ==> Q |] \
       
    51 \  ==> Q";
       
    52 by (cut_facts_tac [major RS (intrelE_lemma RS mp)] 1);
       
    53 by (REPEAT (eresolve_tac [asm_rl,exE,conjE,minor] 1));
       
    54 qed "intrelE";
       
    55 
       
    56 val intrel_cs = rel_cs addSIs [intrelI] addSEs [intrelE];
       
    57 
       
    58 goal Integ.thy "<<x1,y1>,<x2,y2>>: intrel = (x1+y2 = x2+y1)";
       
    59 by (fast_tac intrel_cs 1);
       
    60 qed "intrel_iff";
       
    61 
       
    62 goal Integ.thy "<x,x>: intrel";
       
    63 by (rtac (surjective_pairing RS ssubst) 1 THEN rtac (refl RS intrelI) 1);
       
    64 qed "intrel_refl";
       
    65 
       
    66 goalw Integ.thy [equiv_def, refl_def, sym_def, trans_def]
       
    67     "equiv {x::(nat*nat).True} intrel";
       
    68 by (fast_tac (intrel_cs addSIs [intrel_refl] 
       
    69                         addSEs [sym, integ_trans_lemma]) 1);
       
    70 qed "equiv_intrel";
       
    71 
       
    72 val equiv_intrel_iff =
       
    73     [TrueI, TrueI] MRS 
       
    74     ([CollectI, CollectI] MRS 
       
    75     (equiv_intrel RS eq_equiv_class_iff));
       
    76 
       
    77 goalw Integ.thy  [Integ_def,intrel_def,quotient_def] "intrel^^{<x,y>}:Integ";
       
    78 by (fast_tac set_cs 1);
       
    79 qed "intrel_in_integ";
       
    80 
       
    81 goal Integ.thy "inj_onto Abs_Integ Integ";
       
    82 by (rtac inj_onto_inverseI 1);
       
    83 by (etac Abs_Integ_inverse 1);
       
    84 qed "inj_onto_Abs_Integ";
       
    85 
       
    86 val intrel_ss = 
       
    87     arith_ss addsimps [equiv_intrel_iff, inj_onto_Abs_Integ RS inj_onto_iff,
       
    88 		       intrel_iff, intrel_in_integ, Abs_Integ_inverse];
       
    89 
       
    90 goal Integ.thy "inj(Rep_Integ)";
       
    91 by (rtac inj_inverseI 1);
       
    92 by (rtac Rep_Integ_inverse 1);
       
    93 qed "inj_Rep_Integ";
       
    94 
       
    95 
       
    96 
       
    97 
       
    98 (** znat: the injection from nat to Integ **)
       
    99 
       
   100 goal Integ.thy "inj(znat)";
       
   101 by (rtac injI 1);
       
   102 by (rewtac znat_def);
       
   103 by (dtac (inj_onto_Abs_Integ RS inj_ontoD) 1);
       
   104 by (REPEAT (rtac intrel_in_integ 1));
       
   105 by (dtac eq_equiv_class 1);
       
   106 by (rtac equiv_intrel 1);
       
   107 by (fast_tac set_cs 1);
       
   108 by (safe_tac intrel_cs);
       
   109 by (asm_full_simp_tac arith_ss 1);
       
   110 qed "inj_znat";
       
   111 
       
   112 
       
   113 (**** zminus: unary negation on Integ ****)
       
   114 
       
   115 goalw Integ.thy [congruent_def]
       
   116   "congruent intrel (%p. split (%x y. intrel^^{<y,x>}) p)";
       
   117 by (safe_tac intrel_cs);
       
   118 by (asm_simp_tac (intrel_ss addsimps add_ac) 1);
       
   119 qed "zminus_congruent";
       
   120 
       
   121 
       
   122 (*Resolve th against the corresponding facts for zminus*)
       
   123 val zminus_ize = RSLIST [equiv_intrel, zminus_congruent];
       
   124 
       
   125 goalw Integ.thy [zminus_def]
       
   126       "$~ Abs_Integ(intrel^^{<x,y>}) = Abs_Integ(intrel ^^ {<y,x>})";
       
   127 by (res_inst_tac [("f","Abs_Integ")] arg_cong 1);
       
   128 by (simp_tac (set_ss addsimps 
       
   129    [intrel_in_integ RS Abs_Integ_inverse,zminus_ize UN_equiv_class]) 1);
       
   130 by (rewtac split_def);
       
   131 by (simp_tac prod_ss 1);
       
   132 qed "zminus";
       
   133 
       
   134 (*by lcp*)
       
   135 val [prem] = goal Integ.thy
       
   136     "(!!x y. z = Abs_Integ(intrel^^{<x,y>}) ==> P) ==> P";
       
   137 by (res_inst_tac [("x1","z")] 
       
   138     (rewrite_rule [Integ_def] Rep_Integ RS quotientE) 1);
       
   139 by (dres_inst_tac [("f","Abs_Integ")] arg_cong 1);
       
   140 by (res_inst_tac [("p","x")] PairE 1);
       
   141 by (rtac prem 1);
       
   142 by (asm_full_simp_tac (HOL_ss addsimps [Rep_Integ_inverse]) 1);
       
   143 qed "eq_Abs_Integ";
       
   144 
       
   145 goal Integ.thy "$~ ($~ z) = z";
       
   146 by (res_inst_tac [("z","z")] eq_Abs_Integ 1);
       
   147 by (asm_simp_tac (HOL_ss addsimps [zminus]) 1);
       
   148 qed "zminus_zminus";
       
   149 
       
   150 goal Integ.thy "inj(zminus)";
       
   151 by (rtac injI 1);
       
   152 by (dres_inst_tac [("f","zminus")] arg_cong 1);
       
   153 by (asm_full_simp_tac (HOL_ss addsimps [zminus_zminus]) 1);
       
   154 qed "inj_zminus";
       
   155 
       
   156 goalw Integ.thy [znat_def] "$~ ($#0) = $#0";
       
   157 by (simp_tac (arith_ss addsimps [zminus]) 1);
       
   158 qed "zminus_0";
       
   159 
       
   160 
       
   161 (**** znegative: the test for negative integers ****)
       
   162 
       
   163 goal Arith.thy "!!m x n::nat. n+m=x ==> m<=x";
       
   164 by (dtac (disjI2 RS less_or_eq_imp_le) 1);
       
   165 by (asm_full_simp_tac (arith_ss addsimps add_ac) 1);
       
   166 by (dtac add_leD1 1);
       
   167 by (assume_tac 1);
       
   168 qed "not_znegative_znat_lemma";
       
   169 
       
   170 
       
   171 goalw Integ.thy [znegative_def, znat_def]
       
   172     "~ znegative($# n)";
       
   173 by (simp_tac intrel_ss 1);
       
   174 by (safe_tac intrel_cs);
       
   175 by (rtac ccontr 1);
       
   176 by (etac notE 1);
       
   177 by (asm_full_simp_tac arith_ss 1);
       
   178 by (dtac not_znegative_znat_lemma 1);
       
   179 by (fast_tac (HOL_cs addDs [leD]) 1);
       
   180 qed "not_znegative_znat";
       
   181 
       
   182 goalw Integ.thy [znegative_def, znat_def] "znegative($~ $# Suc(n))";
       
   183 by (simp_tac (intrel_ss addsimps [zminus]) 1);
       
   184 by (REPEAT (ares_tac [exI, conjI] 1));
       
   185 by (rtac (intrelI RS ImageI) 2);
       
   186 by (rtac singletonI 3);
       
   187 by (simp_tac arith_ss 2);
       
   188 by (rtac less_add_Suc1 1);
       
   189 qed "znegative_zminus_znat";
       
   190 
       
   191 
       
   192 (**** zmagnitude: magnitide of an integer, as a natural number ****)
       
   193 
       
   194 goal Arith.thy "!!n::nat. n - Suc(n+m)=0";
       
   195 by (nat_ind_tac "n" 1);
       
   196 by (ALLGOALS(asm_simp_tac arith_ss));
       
   197 qed "diff_Suc_add_0";
       
   198 
       
   199 goal Arith.thy "Suc((n::nat)+m)-n=Suc(m)";
       
   200 by (nat_ind_tac "n" 1);
       
   201 by (ALLGOALS(asm_simp_tac arith_ss));
       
   202 qed "diff_Suc_add_inverse";
       
   203 
       
   204 goalw Integ.thy [congruent_def]
       
   205     "congruent intrel (split (%x y. intrel^^{<(y-x) + (x-(y::nat)),0>}))";
       
   206 by (safe_tac intrel_cs);
       
   207 by (asm_simp_tac intrel_ss 1);
       
   208 by (etac rev_mp 1);
       
   209 by (res_inst_tac [("m","x1"),("n","y1")] diff_induct 1);
       
   210 by (asm_simp_tac (arith_ss addsimps [inj_Suc RS inj_eq]) 3);
       
   211 by (asm_simp_tac (arith_ss addsimps [diff_add_inverse,diff_add_0]) 2);
       
   212 by (asm_simp_tac arith_ss 1);
       
   213 by (rtac impI 1);
       
   214 by (etac subst 1);
       
   215 by (res_inst_tac [("m1","x")] (add_commute RS ssubst) 1);
       
   216 by (asm_simp_tac (arith_ss addsimps [diff_add_inverse,diff_add_0]) 1);
       
   217 by (rtac impI 1);
       
   218 by (asm_simp_tac (arith_ss addsimps
       
   219 		  [diff_add_inverse, diff_add_0, diff_Suc_add_0,
       
   220 		   diff_Suc_add_inverse]) 1);
       
   221 qed "zmagnitude_congruent";
       
   222 
       
   223 (*Resolve th against the corresponding facts for zmagnitude*)
       
   224 val zmagnitude_ize = RSLIST [equiv_intrel, zmagnitude_congruent];
       
   225 
       
   226 
       
   227 goalw Integ.thy [zmagnitude_def]
       
   228     "zmagnitude (Abs_Integ(intrel^^{<x,y>})) = \
       
   229 \    Abs_Integ(intrel^^{<(y - x) + (x - y),0>})";
       
   230 by (res_inst_tac [("f","Abs_Integ")] arg_cong 1);
       
   231 by (asm_simp_tac (intrel_ss addsimps [zmagnitude_ize UN_equiv_class]) 1);
       
   232 qed "zmagnitude";
       
   233 
       
   234 goalw Integ.thy [znat_def] "zmagnitude($# n) = $#n";
       
   235 by (asm_simp_tac (intrel_ss addsimps [zmagnitude]) 1);
       
   236 qed "zmagnitude_znat";
       
   237 
       
   238 goalw Integ.thy [znat_def] "zmagnitude($~ $# n) = $#n";
       
   239 by (asm_simp_tac (intrel_ss addsimps [zmagnitude, zminus]) 1);
       
   240 qed "zmagnitude_zminus_znat";
       
   241 
       
   242 
       
   243 (**** zadd: addition on Integ ****)
       
   244 
       
   245 (** Congruence property for addition **)
       
   246 
       
   247 goalw Integ.thy [congruent2_def]
       
   248     "congruent2 intrel (%p1 p2.                  \
       
   249 \         split (%x1 y1. split (%x2 y2. intrel^^{<x1+x2, y1+y2>}) p2) p1)";
       
   250 (*Proof via congruent2_commuteI seems longer*)
       
   251 by (safe_tac intrel_cs);
       
   252 by (asm_simp_tac (intrel_ss addsimps [add_assoc]) 1);
       
   253 (*The rest should be trivial, but rearranging terms is hard*)
       
   254 by (res_inst_tac [("x1","x1a")] (add_left_commute RS ssubst) 1);
       
   255 by (asm_simp_tac (arith_ss addsimps [add_assoc RS sym]) 1);
       
   256 by (asm_simp_tac (arith_ss addsimps add_ac) 1);
       
   257 qed "zadd_congruent2";
       
   258 
       
   259 (*Resolve th against the corresponding facts for zadd*)
       
   260 val zadd_ize = RSLIST [equiv_intrel, zadd_congruent2];
       
   261 
       
   262 goalw Integ.thy [zadd_def]
       
   263   "Abs_Integ(intrel^^{<x1,y1>}) + Abs_Integ(intrel^^{<x2,y2>}) = \
       
   264 \  Abs_Integ(intrel^^{<x1+x2, y1+y2>})";
       
   265 by (asm_simp_tac
       
   266     (intrel_ss addsimps [zadd_ize UN_equiv_class2]) 1);
       
   267 qed "zadd";
       
   268 
       
   269 goalw Integ.thy [znat_def] "$#0 + z = z";
       
   270 by (res_inst_tac [("z","z")] eq_Abs_Integ 1);
       
   271 by (asm_simp_tac (arith_ss addsimps [zadd]) 1);
       
   272 qed "zadd_0";
       
   273 
       
   274 goal Integ.thy "$~ (z + w) = $~ z + $~ w";
       
   275 by (res_inst_tac [("z","z")] eq_Abs_Integ 1);
       
   276 by (res_inst_tac [("z","w")] eq_Abs_Integ 1);
       
   277 by (asm_simp_tac (arith_ss addsimps [zminus,zadd]) 1);
       
   278 qed "zminus_zadd_distrib";
       
   279 
       
   280 goal Integ.thy "(z::int) + w = w + z";
       
   281 by (res_inst_tac [("z","z")] eq_Abs_Integ 1);
       
   282 by (res_inst_tac [("z","w")] eq_Abs_Integ 1);
       
   283 by (asm_simp_tac (intrel_ss addsimps (add_ac @ [zadd])) 1);
       
   284 qed "zadd_commute";
       
   285 
       
   286 goal Integ.thy "((z1::int) + z2) + z3 = z1 + (z2 + z3)";
       
   287 by (res_inst_tac [("z","z1")] eq_Abs_Integ 1);
       
   288 by (res_inst_tac [("z","z2")] eq_Abs_Integ 1);
       
   289 by (res_inst_tac [("z","z3")] eq_Abs_Integ 1);
       
   290 by (asm_simp_tac (arith_ss addsimps [zadd, add_assoc]) 1);
       
   291 qed "zadd_assoc";
       
   292 
       
   293 (*For AC rewriting*)
       
   294 goal Integ.thy "(x::int)+(y+z)=y+(x+z)";
       
   295 by (rtac (zadd_commute RS trans) 1);
       
   296 by (rtac (zadd_assoc RS trans) 1);
       
   297 by (rtac (zadd_commute RS arg_cong) 1);
       
   298 qed "zadd_left_commute";
       
   299 
       
   300 (*Integer addition is an AC operator*)
       
   301 val zadd_ac = [zadd_assoc,zadd_commute,zadd_left_commute];
       
   302 
       
   303 goalw Integ.thy [znat_def] "$# (m + n) = ($#m) + ($#n)";
       
   304 by (asm_simp_tac (arith_ss addsimps [zadd]) 1);
       
   305 qed "znat_add";
       
   306 
       
   307 goalw Integ.thy [znat_def] "z + ($~ z) = $#0";
       
   308 by (res_inst_tac [("z","z")] eq_Abs_Integ 1);
       
   309 by (asm_simp_tac (intrel_ss addsimps [zminus, zadd, add_commute]) 1);
       
   310 qed "zadd_zminus_inverse";
       
   311 
       
   312 goal Integ.thy "($~ z) + z = $#0";
       
   313 by (rtac (zadd_commute RS trans) 1);
       
   314 by (rtac zadd_zminus_inverse 1);
       
   315 qed "zadd_zminus_inverse2";
       
   316 
       
   317 goal Integ.thy "z + $#0 = z";
       
   318 by (rtac (zadd_commute RS trans) 1);
       
   319 by (rtac zadd_0 1);
       
   320 qed "zadd_0_right";
       
   321 
       
   322 
       
   323 (*Need properties of subtraction?  Or use $- just as an abbreviation!*)
       
   324 
       
   325 (**** zmult: multiplication on Integ ****)
       
   326 
       
   327 (** Congruence property for multiplication **)
       
   328 
       
   329 goal Integ.thy "((k::nat) + l) + (m + n) = (k + m) + (n + l)";
       
   330 by (simp_tac (arith_ss addsimps add_ac) 1);
       
   331 qed "zmult_congruent_lemma";
       
   332 
       
   333 goal Integ.thy 
       
   334     "congruent2 intrel (%p1 p2.  		\
       
   335 \               split (%x1 y1. split (%x2 y2. 	\
       
   336 \                   intrel^^{<x1*x2 + y1*y2, x1*y2 + y1*x2>}) p2) p1)";
       
   337 by (rtac (equiv_intrel RS congruent2_commuteI) 1);
       
   338 by (safe_tac intrel_cs);
       
   339 by (rewtac split_def);
       
   340 by (simp_tac (arith_ss addsimps add_ac@mult_ac) 1);
       
   341 by (asm_simp_tac (arith_ss addsimps add_ac@mult_ac) 1);
       
   342 by (rtac (intrelI RS(equiv_intrel RS equiv_class_eq)) 1);
       
   343 by (rtac (zmult_congruent_lemma RS trans) 1);
       
   344 by (rtac (zmult_congruent_lemma RS trans RS sym) 1);
       
   345 by (rtac (zmult_congruent_lemma RS trans RS sym) 1);
       
   346 by (rtac (zmult_congruent_lemma RS trans RS sym) 1);
       
   347 by (asm_simp_tac (HOL_ss addsimps [add_mult_distrib RS sym]) 1);
       
   348 by (asm_simp_tac (arith_ss addsimps add_ac@mult_ac) 1);
       
   349 qed "zmult_congruent2";
       
   350 
       
   351 (*Resolve th against the corresponding facts for zmult*)
       
   352 val zmult_ize = RSLIST [equiv_intrel, zmult_congruent2];
       
   353 
       
   354 goalw Integ.thy [zmult_def]
       
   355    "Abs_Integ((intrel^^{<x1,y1>})) * Abs_Integ((intrel^^{<x2,y2>})) = 	\
       
   356 \   Abs_Integ(intrel ^^ {<x1*x2 + y1*y2, x1*y2 + y1*x2>})";
       
   357 by (simp_tac (intrel_ss addsimps [zmult_ize UN_equiv_class2]) 1);
       
   358 qed "zmult";
       
   359 
       
   360 goalw Integ.thy [znat_def] "$#0 * z = $#0";
       
   361 by (res_inst_tac [("z","z")] eq_Abs_Integ 1);
       
   362 by (asm_simp_tac (arith_ss addsimps [zmult]) 1);
       
   363 qed "zmult_0";
       
   364 
       
   365 goalw Integ.thy [znat_def] "$#Suc(0) * z = z";
       
   366 by (res_inst_tac [("z","z")] eq_Abs_Integ 1);
       
   367 by (asm_simp_tac (arith_ss addsimps [zmult, add_0_right]) 1);
       
   368 qed "zmult_1";
       
   369 
       
   370 goal Integ.thy "($~ z) * w = $~ (z * w)";
       
   371 by (res_inst_tac [("z","z")] eq_Abs_Integ 1);
       
   372 by (res_inst_tac [("z","w")] eq_Abs_Integ 1);
       
   373 by (asm_simp_tac (intrel_ss addsimps ([zminus, zmult] @ add_ac)) 1);
       
   374 qed "zmult_zminus";
       
   375 
       
   376 
       
   377 goal Integ.thy "($~ z) * ($~ w) = (z * w)";
       
   378 by (res_inst_tac [("z","z")] eq_Abs_Integ 1);
       
   379 by (res_inst_tac [("z","w")] eq_Abs_Integ 1);
       
   380 by (asm_simp_tac (intrel_ss addsimps ([zminus, zmult] @ add_ac)) 1);
       
   381 qed "zmult_zminus_zminus";
       
   382 
       
   383 goal Integ.thy "(z::int) * w = w * z";
       
   384 by (res_inst_tac [("z","z")] eq_Abs_Integ 1);
       
   385 by (res_inst_tac [("z","w")] eq_Abs_Integ 1);
       
   386 by (asm_simp_tac (intrel_ss addsimps ([zmult] @ add_ac @ mult_ac)) 1);
       
   387 qed "zmult_commute";
       
   388 
       
   389 goal Integ.thy "z * $# 0 = $#0";
       
   390 by (rtac ([zmult_commute, zmult_0] MRS trans) 1);
       
   391 qed "zmult_0_right";
       
   392 
       
   393 goal Integ.thy "z * $#Suc(0) = z";
       
   394 by (rtac ([zmult_commute, zmult_1] MRS trans) 1);
       
   395 qed "zmult_1_right";
       
   396 
       
   397 goal Integ.thy "((z1::int) * z2) * z3 = z1 * (z2 * z3)";
       
   398 by (res_inst_tac [("z","z1")] eq_Abs_Integ 1);
       
   399 by (res_inst_tac [("z","z2")] eq_Abs_Integ 1);
       
   400 by (res_inst_tac [("z","z3")] eq_Abs_Integ 1);
       
   401 by (asm_simp_tac (intrel_ss addsimps ([zmult] @ add_ac @ mult_ac)) 1);
       
   402 qed "zmult_assoc";
       
   403 
       
   404 (*For AC rewriting*)
       
   405 qed_goal "zmult_left_commute" Integ.thy
       
   406     "(z1::int)*(z2*z3) = z2*(z1*z3)"
       
   407  (fn _ => [rtac (zmult_commute RS trans) 1, rtac (zmult_assoc RS trans) 1,
       
   408            rtac (zmult_commute RS arg_cong) 1]);
       
   409 
       
   410 (*Integer multiplication is an AC operator*)
       
   411 val zmult_ac = [zmult_assoc, zmult_commute, zmult_left_commute];
       
   412 
       
   413 goal Integ.thy "((z1::int) + z2) * w = (z1 * w) + (z2 * w)";
       
   414 by (res_inst_tac [("z","z1")] eq_Abs_Integ 1);
       
   415 by (res_inst_tac [("z","z2")] eq_Abs_Integ 1);
       
   416 by (res_inst_tac [("z","w")] eq_Abs_Integ 1);
       
   417 by (asm_simp_tac 
       
   418     (intrel_ss addsimps ([zadd, zmult, add_mult_distrib] @ 
       
   419 			 add_ac @ mult_ac)) 1);
       
   420 qed "zadd_zmult_distrib";
       
   421 
       
   422 val zmult_commute'= read_instantiate [("z","w")] zmult_commute;
       
   423 
       
   424 goal Integ.thy "w * ($~ z) = $~ (w * z)";
       
   425 by (simp_tac (HOL_ss addsimps [zmult_commute', zmult_zminus]) 1);
       
   426 qed "zmult_zminus_right";
       
   427 
       
   428 goal Integ.thy "(w::int) * (z1 + z2) = (w * z1) + (w * z2)";
       
   429 by (simp_tac (HOL_ss addsimps [zmult_commute',zadd_zmult_distrib]) 1);
       
   430 qed "zadd_zmult_distrib2";
       
   431 
       
   432 val zadd_simps = 
       
   433     [zadd_0, zadd_0_right, zadd_zminus_inverse, zadd_zminus_inverse2];
       
   434 
       
   435 val zminus_simps = [zminus_zminus, zminus_0, zminus_zadd_distrib];
       
   436 
       
   437 val zmult_simps = [zmult_0, zmult_1, zmult_0_right, zmult_1_right, 
       
   438 		   zmult_zminus, zmult_zminus_right];
       
   439 
       
   440 val integ_ss =
       
   441     arith_ss addsimps (zadd_simps @ zminus_simps @ zmult_simps @ 
       
   442 		       [zmagnitude_znat, zmagnitude_zminus_znat]);
       
   443 
       
   444 
       
   445 (**** Additional Theorems (by Mattolini; proofs mainly by lcp) ****)
       
   446 
       
   447 (* Some Theorems about zsuc and zpred *)
       
   448 goalw Integ.thy [zsuc_def] "$#(Suc(n)) = zsuc($# n)";
       
   449 by (simp_tac (arith_ss addsimps [znat_add RS sym]) 1);
       
   450 qed "znat_Suc";
       
   451 
       
   452 goalw Integ.thy [zpred_def,zsuc_def,zdiff_def] "$~ zsuc(z) = zpred($~ z)";
       
   453 by (simp_tac integ_ss 1);
       
   454 qed "zminus_zsuc";
       
   455 
       
   456 goalw Integ.thy [zpred_def,zsuc_def,zdiff_def] "$~ zpred(z) = zsuc($~ z)";
       
   457 by (simp_tac integ_ss 1);
       
   458 qed "zminus_zpred";
       
   459 
       
   460 goalw Integ.thy [zsuc_def,zpred_def,zdiff_def]
       
   461    "zpred(zsuc(z)) = z";
       
   462 by (simp_tac (integ_ss addsimps [zadd_assoc]) 1);
       
   463 qed "zpred_zsuc";
       
   464 
       
   465 goalw Integ.thy [zsuc_def,zpred_def,zdiff_def]
       
   466    "zsuc(zpred(z)) = z";
       
   467 by (simp_tac (integ_ss addsimps [zadd_assoc]) 1);
       
   468 qed "zsuc_zpred";
       
   469 
       
   470 goal Integ.thy "(zpred(z)=w) = (z=zsuc(w))";
       
   471 by (safe_tac HOL_cs);
       
   472 by (rtac (zsuc_zpred RS sym) 1);
       
   473 by (rtac zpred_zsuc 1);
       
   474 qed "zpred_to_zsuc";
       
   475 
       
   476 goal Integ.thy "(zsuc(z)=w)=(z=zpred(w))";
       
   477 by (safe_tac HOL_cs);
       
   478 by (rtac (zpred_zsuc RS sym) 1);
       
   479 by (rtac zsuc_zpred 1);
       
   480 qed "zsuc_to_zpred";
       
   481 
       
   482 goal Integ.thy "($~ z = w) = (z = $~ w)";
       
   483 by (safe_tac HOL_cs);
       
   484 by (rtac (zminus_zminus RS sym) 1);
       
   485 by (rtac zminus_zminus 1);
       
   486 qed "zminus_exchange";
       
   487 
       
   488 goal Integ.thy"(zsuc(z)=zsuc(w)) = (z=w)";
       
   489 by (safe_tac intrel_cs);
       
   490 by (dres_inst_tac [("f","zpred")] arg_cong 1);
       
   491 by (asm_full_simp_tac (HOL_ss addsimps [zpred_zsuc]) 1);
       
   492 qed "bijective_zsuc";
       
   493 
       
   494 goal Integ.thy"(zpred(z)=zpred(w)) = (z=w)";
       
   495 by (safe_tac intrel_cs);
       
   496 by (dres_inst_tac [("f","zsuc")] arg_cong 1);
       
   497 by (asm_full_simp_tac (HOL_ss addsimps [zsuc_zpred]) 1);
       
   498 qed "bijective_zpred";
       
   499 
       
   500 (* Additional Theorems about zadd *)
       
   501 
       
   502 goalw Integ.thy [zsuc_def] "zsuc(z) + w = zsuc(z+w)";
       
   503 by (simp_tac (arith_ss addsimps zadd_ac) 1);
       
   504 qed "zadd_zsuc";
       
   505 
       
   506 goalw Integ.thy [zsuc_def] "w + zsuc(z) = zsuc(w+z)";
       
   507 by (simp_tac (arith_ss addsimps zadd_ac) 1);
       
   508 qed "zadd_zsuc_right";
       
   509 
       
   510 goalw Integ.thy [zpred_def,zdiff_def] "zpred(z) + w = zpred(z+w)";
       
   511 by (simp_tac (arith_ss addsimps zadd_ac) 1);
       
   512 qed "zadd_zpred";
       
   513 
       
   514 goalw Integ.thy [zpred_def,zdiff_def] "w + zpred(z) = zpred(w+z)";
       
   515 by (simp_tac (arith_ss addsimps zadd_ac) 1);
       
   516 qed "zadd_zpred_right";
       
   517 
       
   518 
       
   519 (* Additional Theorems about zmult *)
       
   520 
       
   521 goalw Integ.thy [zsuc_def] "zsuc(w) * z = z + w * z";
       
   522 by (simp_tac (integ_ss addsimps [zadd_zmult_distrib, zadd_commute]) 1);
       
   523 qed "zmult_zsuc";
       
   524 
       
   525 goalw Integ.thy [zsuc_def] "z * zsuc(w) = z + w * z";
       
   526 by (simp_tac 
       
   527     (integ_ss addsimps [zadd_zmult_distrib2, zadd_commute, zmult_commute]) 1);
       
   528 qed "zmult_zsuc_right";
       
   529 
       
   530 goalw Integ.thy [zpred_def, zdiff_def] "zpred(w) * z = w * z - z";
       
   531 by (simp_tac (integ_ss addsimps [zadd_zmult_distrib]) 1);
       
   532 qed "zmult_zpred";
       
   533 
       
   534 goalw Integ.thy [zpred_def, zdiff_def] "z * zpred(w) = w * z - z";
       
   535 by (simp_tac (integ_ss addsimps [zadd_zmult_distrib2, zmult_commute]) 1);
       
   536 qed "zmult_zpred_right";
       
   537 
       
   538 (* Further Theorems about zsuc and zpred *)
       
   539 goal Integ.thy "$#Suc(m) ~= $#0";
       
   540 by (simp_tac (integ_ss addsimps [inj_znat RS inj_eq]) 1);
       
   541 qed "znat_Suc_not_znat_Zero";
       
   542 
       
   543 bind_thm ("znat_Zero_not_znat_Suc", (znat_Suc_not_znat_Zero RS not_sym));
       
   544 
       
   545 
       
   546 goalw Integ.thy [zsuc_def,znat_def] "w ~= zsuc(w)";
       
   547 by (res_inst_tac [("z","w")] eq_Abs_Integ 1);
       
   548 by (asm_full_simp_tac (intrel_ss addsimps [zadd]) 1);
       
   549 qed "n_not_zsuc_n";
       
   550 
       
   551 val zsuc_n_not_n = n_not_zsuc_n RS not_sym;
       
   552 
       
   553 goal Integ.thy "w ~= zpred(w)";
       
   554 by (safe_tac HOL_cs);
       
   555 by (dres_inst_tac [("x","w"),("f","zsuc")] arg_cong 1);
       
   556 by (asm_full_simp_tac (HOL_ss addsimps [zsuc_zpred,zsuc_n_not_n]) 1);
       
   557 qed "n_not_zpred_n";
       
   558 
       
   559 val zpred_n_not_n = n_not_zpred_n RS not_sym;
       
   560 
       
   561 
       
   562 (* Theorems about less and less_equal *)
       
   563 
       
   564 goalw Integ.thy [zless_def, znegative_def, zdiff_def, znat_def] 
       
   565     "!!w. w<z ==> ? n. z = w + $#(Suc(n))";
       
   566 by (res_inst_tac [("z","z")] eq_Abs_Integ 1);
       
   567 by (res_inst_tac [("z","w")] eq_Abs_Integ 1);
       
   568 by (safe_tac intrel_cs);
       
   569 by (asm_full_simp_tac (intrel_ss addsimps [zadd, zminus]) 1);
       
   570 by (safe_tac (intrel_cs addSDs [less_eq_Suc_add]));
       
   571 by (res_inst_tac [("x","k")] exI 1);
       
   572 by (asm_full_simp_tac (HOL_ss addsimps ([add_Suc RS sym] @ add_ac)) 1);
       
   573 (*To cancel x2, rename it to be first!*)
       
   574 by (rename_tac "a b c" 1);
       
   575 by (asm_full_simp_tac (HOL_ss addsimps (add_left_cancel::add_ac)) 1);
       
   576 qed "zless_eq_zadd_Suc";
       
   577 
       
   578 goalw Integ.thy [zless_def, znegative_def, zdiff_def, znat_def] 
       
   579     "z < z + $#(Suc(n))";
       
   580 by (res_inst_tac [("z","z")] eq_Abs_Integ 1);
       
   581 by (safe_tac intrel_cs);
       
   582 by (simp_tac (intrel_ss addsimps [zadd, zminus]) 1);
       
   583 by (REPEAT_SOME (ares_tac [refl, exI, singletonI, ImageI, conjI, intrelI]));
       
   584 by (rtac le_less_trans 1);
       
   585 by (rtac lessI 2);
       
   586 by (asm_simp_tac (arith_ss addsimps ([le_add1,add_left_cancel_le]@add_ac)) 1);
       
   587 qed "zless_zadd_Suc";
       
   588 
       
   589 goal Integ.thy "!!z1 z2 z3. [| z1<z2; z2<z3 |] ==> z1 < (z3::int)";
       
   590 by (safe_tac (HOL_cs addSDs [zless_eq_zadd_Suc]));
       
   591 by (simp_tac 
       
   592     (arith_ss addsimps [zadd_assoc, zless_zadd_Suc, znat_add RS sym]) 1);
       
   593 qed "zless_trans";
       
   594 
       
   595 goalw Integ.thy [zsuc_def] "z<zsuc(z)";
       
   596 by (rtac zless_zadd_Suc 1);
       
   597 qed "zlessI";
       
   598 
       
   599 val zless_zsucI = zlessI RSN (2,zless_trans);
       
   600 
       
   601 goal Integ.thy "!!z w::int. z<w ==> ~w<z";
       
   602 by (safe_tac (HOL_cs addSDs [zless_eq_zadd_Suc]));
       
   603 by (res_inst_tac [("z","z")] eq_Abs_Integ 1);
       
   604 by (safe_tac intrel_cs);
       
   605 by (asm_full_simp_tac (intrel_ss addsimps ([znat_def, zadd])) 1);
       
   606 by (asm_full_simp_tac
       
   607  (HOL_ss addsimps [add_left_cancel, add_assoc, add_Suc_right RS sym]) 1);
       
   608 by (resolve_tac [less_not_refl2 RS notE] 1);
       
   609 by (etac sym 2);
       
   610 by (REPEAT (resolve_tac [lessI, trans_less_add2, less_SucI] 1));
       
   611 qed "zless_not_sym";
       
   612 
       
   613 (* [| n<m; m<n |] ==> R *)
       
   614 bind_thm ("zless_asym", (zless_not_sym RS notE));
       
   615 
       
   616 goal Integ.thy "!!z::int. ~ z<z";
       
   617 by (resolve_tac [zless_asym RS notI] 1);
       
   618 by (REPEAT (assume_tac 1));
       
   619 qed "zless_not_refl";
       
   620 
       
   621 (* z<z ==> R *)
       
   622 bind_thm ("zless_anti_refl", (zless_not_refl RS notE));
       
   623 
       
   624 goal Integ.thy "!!w. z<w ==> w ~= (z::int)";
       
   625 by(fast_tac (HOL_cs addEs [zless_anti_refl]) 1);
       
   626 qed "zless_not_refl2";
       
   627 
       
   628 
       
   629 (*"Less than" is a linear ordering*)
       
   630 goalw Integ.thy [zless_def, znegative_def, zdiff_def] 
       
   631     "z<w | z=w | w<(z::int)";
       
   632 by (res_inst_tac [("z","z")] eq_Abs_Integ 1);
       
   633 by (res_inst_tac [("z","w")] eq_Abs_Integ 1);
       
   634 by (safe_tac intrel_cs);
       
   635 by (asm_full_simp_tac
       
   636     (intrel_ss addsimps [zadd, zminus, Image_iff, Bex_def]) 1);
       
   637 by (res_inst_tac [("m1", "x+ya"), ("n1", "xa+y")] (less_linear RS disjE) 1);
       
   638 by (etac disjE 2);
       
   639 by (assume_tac 2);
       
   640 by (DEPTH_SOLVE
       
   641     (swap_res_tac [exI] 1 THEN 
       
   642      swap_res_tac [exI] 1 THEN 
       
   643      etac conjI 1 THEN 
       
   644      simp_tac (arith_ss addsimps add_ac)  1));
       
   645 qed "zless_linear";
       
   646 
       
   647 
       
   648 (*** Properties of <= ***)
       
   649 
       
   650 goalw Integ.thy  [zless_def, znegative_def, zdiff_def, znat_def]
       
   651     "($#m < $#n) = (m<n)";
       
   652 by (simp_tac
       
   653     (intrel_ss addsimps [zadd, zminus, Image_iff, Bex_def]) 1);
       
   654 by (fast_tac (HOL_cs addIs [add_commute] addSEs [less_add_eq_less]) 1);
       
   655 qed "zless_eq_less";
       
   656 
       
   657 goalw Integ.thy [zle_def, le_def] "($#m <= $#n) = (m<=n)";
       
   658 by (simp_tac (HOL_ss addsimps [zless_eq_less]) 1);
       
   659 qed "zle_eq_le";
       
   660 
       
   661 goalw Integ.thy [zle_def] "!!w. ~(w<z) ==> z<=(w::int)";
       
   662 by (assume_tac 1);
       
   663 qed "zleI";
       
   664 
       
   665 goalw Integ.thy [zle_def] "!!w. z<=w ==> ~(w<(z::int))";
       
   666 by (assume_tac 1);
       
   667 qed "zleD";
       
   668 
       
   669 val zleE = make_elim zleD;
       
   670 
       
   671 goalw Integ.thy [zle_def] "!!z. ~ z <= w ==> w<(z::int)";
       
   672 by (fast_tac HOL_cs 1);
       
   673 qed "not_zleE";
       
   674 
       
   675 goalw Integ.thy [zle_def] "!!z. z < w ==> z <= (w::int)";
       
   676 by (fast_tac (HOL_cs addEs [zless_asym]) 1);
       
   677 qed "zless_imp_zle";
       
   678 
       
   679 goalw Integ.thy [zle_def] "!!z. z <= w ==> z < w | z=(w::int)";
       
   680 by (cut_facts_tac [zless_linear] 1);
       
   681 by (fast_tac (HOL_cs addEs [zless_anti_refl,zless_asym]) 1);
       
   682 qed "zle_imp_zless_or_eq";
       
   683 
       
   684 goalw Integ.thy [zle_def] "!!z. z<w | z=w ==> z <=(w::int)";
       
   685 by (cut_facts_tac [zless_linear] 1);
       
   686 by (fast_tac (HOL_cs addEs [zless_anti_refl,zless_asym]) 1);
       
   687 qed "zless_or_eq_imp_zle";
       
   688 
       
   689 goal Integ.thy "(x <= (y::int)) = (x < y | x=y)";
       
   690 by (REPEAT(ares_tac [iffI, zless_or_eq_imp_zle, zle_imp_zless_or_eq] 1));
       
   691 qed "zle_eq_zless_or_eq";
       
   692 
       
   693 goal Integ.thy "w <= (w::int)";
       
   694 by (simp_tac (HOL_ss addsimps [zle_eq_zless_or_eq]) 1);
       
   695 qed "zle_refl";
       
   696 
       
   697 val prems = goal Integ.thy "!!i. [| i <= j; j < k |] ==> i < (k::int)";
       
   698 by (dtac zle_imp_zless_or_eq 1);
       
   699 by (fast_tac (HOL_cs addIs [zless_trans]) 1);
       
   700 qed "zle_zless_trans";
       
   701 
       
   702 goal Integ.thy "!!i. [| i <= j; j <= k |] ==> i <= (k::int)";
       
   703 by (EVERY1 [dtac zle_imp_zless_or_eq, dtac zle_imp_zless_or_eq,
       
   704 	    rtac zless_or_eq_imp_zle, fast_tac (HOL_cs addIs [zless_trans])]);
       
   705 qed "zle_trans";
       
   706 
       
   707 goal Integ.thy "!!z. [| z <= w; w <= z |] ==> z = (w::int)";
       
   708 by (EVERY1 [dtac zle_imp_zless_or_eq, dtac zle_imp_zless_or_eq,
       
   709 	    fast_tac (HOL_cs addEs [zless_anti_refl,zless_asym])]);
       
   710 qed "zle_anti_sym";
       
   711 
       
   712 
       
   713 goal Integ.thy "!!w w' z::int. z + w' = z + w ==> w' = w";
       
   714 by (dres_inst_tac [("f", "%x. x + $~z")] arg_cong 1);
       
   715 by (asm_full_simp_tac (integ_ss addsimps zadd_ac) 1);
       
   716 qed "zadd_left_cancel";
       
   717 
       
   718 
       
   719 (*** Monotonicity results ***)
       
   720 
       
   721 goal Integ.thy "!!v w z::int. v < w ==> v + z < w + z";
       
   722 by (safe_tac (HOL_cs addSDs [zless_eq_zadd_Suc]));
       
   723 by (simp_tac (HOL_ss addsimps zadd_ac) 1);
       
   724 by (simp_tac (HOL_ss addsimps [zadd_assoc RS sym, zless_zadd_Suc]) 1);
       
   725 qed "zadd_zless_mono1";
       
   726 
       
   727 goal Integ.thy "!!v w z::int. (v+z < w+z) = (v < w)";
       
   728 by (safe_tac (HOL_cs addSEs [zadd_zless_mono1]));
       
   729 by (dres_inst_tac [("z", "$~z")] zadd_zless_mono1 1);
       
   730 by (asm_full_simp_tac (integ_ss addsimps [zadd_assoc]) 1);
       
   731 qed "zadd_left_cancel_zless";
       
   732 
       
   733 goal Integ.thy "!!v w z::int. (v+z <= w+z) = (v <= w)";
       
   734 by (asm_full_simp_tac
       
   735     (integ_ss addsimps [zle_def, zadd_left_cancel_zless]) 1);
       
   736 qed "zadd_left_cancel_zle";
       
   737 
       
   738 (*"v<=w ==> v+z <= w+z"*)
       
   739 bind_thm ("zadd_zle_mono1", zadd_left_cancel_zle RS iffD2);
       
   740 
       
   741 
       
   742 goal Integ.thy "!!z' z::int. [| w'<=w; z'<=z |] ==> w' + z' <= w + z";
       
   743 by (etac (zadd_zle_mono1 RS zle_trans) 1);
       
   744 by (simp_tac (HOL_ss addsimps [zadd_commute]) 1);
       
   745 (*w moves to the end because it is free while z', z are bound*)
       
   746 by (etac zadd_zle_mono1 1);
       
   747 qed "zadd_zle_mono";
       
   748 
       
   749 goal Integ.thy "!!w z::int. z<=$#0 ==> w+z <= w";
       
   750 by (dres_inst_tac [("z", "w")] zadd_zle_mono1 1);
       
   751 by (asm_full_simp_tac (integ_ss addsimps [zadd_commute]) 1);
       
   752 qed "zadd_zle_self";