| author | krauss | 
| Wed, 28 Apr 2010 17:42:37 +0200 | |
| changeset 36523 | a294e4ebe0a3 | 
| parent 32960 | 69916a850301 | 
| child 45602 | 2a858377c3d2 | 
| permissions | -rw-r--r-- | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32149diff
changeset | 1 | (* Title: ZF/IntDiv_ZF.thy | 
| 26056 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 2 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 3 | Copyright 1999 University of Cambridge | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 4 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 5 | Here is the division algorithm in ML: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 6 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 7 | fun posDivAlg (a,b) = | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 8 | if a<b then (0,a) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 9 | else let val (q,r) = posDivAlg(a, 2*b) | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32149diff
changeset | 10 | in if 0<=r-b then (2*q+1, r-b) else (2*q, r) | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32149diff
changeset | 11 | end | 
| 26056 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 12 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 13 | fun negDivAlg (a,b) = | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 14 | if 0<=a+b then (~1,a+b) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 15 | else let val (q,r) = negDivAlg(a, 2*b) | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32149diff
changeset | 16 | in if 0<=r-b then (2*q+1, r-b) else (2*q, r) | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32149diff
changeset | 17 | end; | 
| 26056 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 18 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 19 | fun negateSnd (q,r:int) = (q,~r); | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 20 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 21 | fun divAlg (a,b) = if 0<=a then | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32149diff
changeset | 22 | if b>0 then posDivAlg (a,b) | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32149diff
changeset | 23 | else if a=0 then (0,0) | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32149diff
changeset | 24 | else negateSnd (negDivAlg (~a,~b)) | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32149diff
changeset | 25 | else | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32149diff
changeset | 26 | if 0<b then negDivAlg (a,b) | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32149diff
changeset | 27 | else negateSnd (posDivAlg (~a,~b)); | 
| 26056 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 28 | *) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 29 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 30 | header{*The Division Operators Div and Mod*}
 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 31 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 32 | theory IntDiv_ZF imports IntArith OrderArith begin | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 33 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 34 | definition | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 35 | quorem :: "[i,i] => o" where | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 36 | "quorem == %<a,b> <q,r>. | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 37 | a = b$*q $+ r & | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 38 | (#0$<b & #0$<=r & r$<b | ~(#0$<b) & b$<r & r $<= #0)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 39 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 40 | definition | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 41 | adjust :: "[i,i] => i" where | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 42 | "adjust(b) == %<q,r>. if #0 $<= r$-b then <#2$*q $+ #1,r$-b> | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 43 | else <#2$*q,r>" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 44 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 45 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 46 | (** the division algorithm **) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 47 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 48 | definition | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 49 | posDivAlg :: "i => i" where | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 50 | (*for the case a>=0, b>0*) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 51 | (*recdef posDivAlg "inv_image less_than (%(a,b). nat_of(a $- b $+ #1))"*) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 52 | "posDivAlg(ab) == | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 53 | wfrec(measure(int*int, %<a,b>. nat_of (a $- b $+ #1)), | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32149diff
changeset | 54 | ab, | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32149diff
changeset | 55 | %<a,b> f. if (a$<b | b$<=#0) then <#0,a> | 
| 26056 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 56 | else adjust(b, f ` <a,#2$*b>))" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 57 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 58 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 59 | (*for the case a<0, b>0*) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 60 | definition | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 61 | negDivAlg :: "i => i" where | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 62 | (*recdef negDivAlg "inv_image less_than (%(a,b). nat_of(- a $- b))"*) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 63 | "negDivAlg(ab) == | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 64 | wfrec(measure(int*int, %<a,b>. nat_of ($- a $- b)), | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32149diff
changeset | 65 | ab, | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32149diff
changeset | 66 | %<a,b> f. if (#0 $<= a$+b | b$<=#0) then <#-1,a$+b> | 
| 26056 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 67 | else adjust(b, f ` <a,#2$*b>))" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 68 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 69 | (*for the general case b\<noteq>0*) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 70 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 71 | definition | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 72 | negateSnd :: "i => i" where | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 73 | "negateSnd == %<q,r>. <q, $-r>" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 74 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 75 | (*The full division algorithm considers all possible signs for a, b | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 76 | including the special case a=0, b<0, because negDivAlg requires a<0*) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 77 | definition | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 78 | divAlg :: "i => i" where | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 79 | "divAlg == | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 80 | %<a,b>. if #0 $<= a then | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 81 | if #0 $<= b then posDivAlg (<a,b>) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 82 | else if a=#0 then <#0,#0> | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 83 | else negateSnd (negDivAlg (<$-a,$-b>)) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 84 | else | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 85 | if #0$<b then negDivAlg (<a,b>) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 86 | else negateSnd (posDivAlg (<$-a,$-b>))" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 87 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 88 | definition | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 89 | zdiv :: "[i,i]=>i" (infixl "zdiv" 70) where | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 90 | "a zdiv b == fst (divAlg (<intify(a), intify(b)>))" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 91 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 92 | definition | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 93 | zmod :: "[i,i]=>i" (infixl "zmod" 70) where | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 94 | "a zmod b == snd (divAlg (<intify(a), intify(b)>))" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 95 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 96 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 97 | (** Some basic laws by Sidi Ehmety (need linear arithmetic!) **) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 98 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 99 | lemma zspos_add_zspos_imp_zspos: "[| #0 $< x; #0 $< y |] ==> #0 $< x $+ y" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 100 | apply (rule_tac y = "y" in zless_trans) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 101 | apply (rule_tac [2] zdiff_zless_iff [THEN iffD1]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 102 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 103 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 104 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 105 | lemma zpos_add_zpos_imp_zpos: "[| #0 $<= x; #0 $<= y |] ==> #0 $<= x $+ y" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 106 | apply (rule_tac y = "y" in zle_trans) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 107 | apply (rule_tac [2] zdiff_zle_iff [THEN iffD1]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 108 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 109 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 110 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 111 | lemma zneg_add_zneg_imp_zneg: "[| x $< #0; y $< #0 |] ==> x $+ y $< #0" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 112 | apply (rule_tac y = "y" in zless_trans) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 113 | apply (rule zless_zdiff_iff [THEN iffD1]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 114 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 115 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 116 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 117 | (* this theorem is used below *) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 118 | lemma zneg_or_0_add_zneg_or_0_imp_zneg_or_0: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 119 | "[| x $<= #0; y $<= #0 |] ==> x $+ y $<= #0" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 120 | apply (rule_tac y = "y" in zle_trans) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 121 | apply (rule zle_zdiff_iff [THEN iffD1]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 122 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 123 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 124 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 125 | lemma zero_lt_zmagnitude: "[| #0 $< k; k \<in> int |] ==> 0 < zmagnitude(k)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 126 | apply (drule zero_zless_imp_znegative_zminus) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 127 | apply (drule_tac [2] zneg_int_of) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 128 | apply (auto simp add: zminus_equation [of k]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 129 | apply (subgoal_tac "0 < zmagnitude ($# succ (n))") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 130 | apply simp | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 131 | apply (simp only: zmagnitude_int_of) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 132 | apply simp | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 133 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 134 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 135 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 136 | (*** Inequality lemmas involving $#succ(m) ***) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 137 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 138 | lemma zless_add_succ_iff: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 139 | "(w $< z $+ $# succ(m)) <-> (w $< z $+ $#m | intify(w) = z $+ $#m)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 140 | apply (auto simp add: zless_iff_succ_zadd zadd_assoc int_of_add [symmetric]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 141 | apply (rule_tac [3] x = "0" in bexI) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 142 | apply (cut_tac m = "m" in int_succ_int_1) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 143 | apply (cut_tac m = "n" in int_succ_int_1) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 144 | apply simp | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 145 | apply (erule natE) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 146 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 147 | apply (rule_tac x = "succ (n) " in bexI) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 148 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 149 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 150 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 151 | lemma zadd_succ_lemma: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 152 | "z \<in> int ==> (w $+ $# succ(m) $<= z) <-> (w $+ $#m $< z)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 153 | apply (simp only: not_zless_iff_zle [THEN iff_sym] zless_add_succ_iff) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 154 | apply (auto intro: zle_anti_sym elim: zless_asym | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 155 | simp add: zless_imp_zle not_zless_iff_zle) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 156 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 157 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 158 | lemma zadd_succ_zle_iff: "(w $+ $# succ(m) $<= z) <-> (w $+ $#m $< z)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 159 | apply (cut_tac z = "intify (z)" in zadd_succ_lemma) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 160 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 161 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 162 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 163 | (** Inequality reasoning **) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 164 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 165 | lemma zless_add1_iff_zle: "(w $< z $+ #1) <-> (w$<=z)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 166 | apply (subgoal_tac "#1 = $# 1") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 167 | apply (simp only: zless_add_succ_iff zle_def) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 168 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 169 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 170 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 171 | lemma add1_zle_iff: "(w $+ #1 $<= z) <-> (w $< z)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 172 | apply (subgoal_tac "#1 = $# 1") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 173 | apply (simp only: zadd_succ_zle_iff) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 174 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 175 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 176 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 177 | lemma add1_left_zle_iff: "(#1 $+ w $<= z) <-> (w $< z)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 178 | apply (subst zadd_commute) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 179 | apply (rule add1_zle_iff) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 180 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 181 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 182 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 183 | (*** Monotonicity of Multiplication ***) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 184 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 185 | lemma zmult_mono_lemma: "k \<in> nat ==> i $<= j ==> i $* $#k $<= j $* $#k" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 186 | apply (induct_tac "k") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 187 | prefer 2 apply (subst int_succ_int_1) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 188 | apply (simp_all (no_asm_simp) add: zadd_zmult_distrib2 zadd_zle_mono) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 189 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 190 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 191 | lemma zmult_zle_mono1: "[| i $<= j; #0 $<= k |] ==> i$*k $<= j$*k" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 192 | apply (subgoal_tac "i $* intify (k) $<= j $* intify (k) ") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 193 | apply (simp (no_asm_use)) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 194 | apply (rule_tac b = "intify (k)" in not_zneg_mag [THEN subst]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 195 | apply (rule_tac [3] zmult_mono_lemma) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 196 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 197 | apply (simp add: znegative_iff_zless_0 not_zless_iff_zle [THEN iff_sym]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 198 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 199 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 200 | lemma zmult_zle_mono1_neg: "[| i $<= j; k $<= #0 |] ==> j$*k $<= i$*k" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 201 | apply (rule zminus_zle_zminus [THEN iffD1]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 202 | apply (simp del: zmult_zminus_right | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 203 | add: zmult_zminus_right [symmetric] zmult_zle_mono1 zle_zminus) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 204 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 205 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 206 | lemma zmult_zle_mono2: "[| i $<= j; #0 $<= k |] ==> k$*i $<= k$*j" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 207 | apply (drule zmult_zle_mono1) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 208 | apply (simp_all add: zmult_commute) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 209 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 210 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 211 | lemma zmult_zle_mono2_neg: "[| i $<= j; k $<= #0 |] ==> k$*j $<= k$*i" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 212 | apply (drule zmult_zle_mono1_neg) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 213 | apply (simp_all add: zmult_commute) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 214 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 215 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 216 | (* $<= monotonicity, BOTH arguments*) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 217 | lemma zmult_zle_mono: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 218 | "[| i $<= j; k $<= l; #0 $<= j; #0 $<= k |] ==> i$*k $<= j$*l" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 219 | apply (erule zmult_zle_mono1 [THEN zle_trans]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 220 | apply assumption | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 221 | apply (erule zmult_zle_mono2) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 222 | apply assumption | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 223 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 224 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 225 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 226 | (** strict, in 1st argument; proof is by induction on k>0 **) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 227 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 228 | lemma zmult_zless_mono2_lemma [rule_format]: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 229 | "[| i$<j; k \<in> nat |] ==> 0<k --> $#k $* i $< $#k $* j" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 230 | apply (induct_tac "k") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 231 | prefer 2 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 232 | apply (subst int_succ_int_1) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 233 | apply (erule natE) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 234 | apply (simp_all add: zadd_zmult_distrib zadd_zless_mono zle_def) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 235 | apply (frule nat_0_le) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 236 | apply (subgoal_tac "i $+ (i $+ $# xa $* i) $< j $+ (j $+ $# xa $* j) ") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 237 | apply (simp (no_asm_use)) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 238 | apply (rule zadd_zless_mono) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 239 | apply (simp_all (no_asm_simp) add: zle_def) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 240 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 241 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 242 | lemma zmult_zless_mono2: "[| i$<j; #0 $< k |] ==> k$*i $< k$*j" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 243 | apply (subgoal_tac "intify (k) $* i $< intify (k) $* j") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 244 | apply (simp (no_asm_use)) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 245 | apply (rule_tac b = "intify (k)" in not_zneg_mag [THEN subst]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 246 | apply (rule_tac [3] zmult_zless_mono2_lemma) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 247 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 248 | apply (simp add: znegative_iff_zless_0) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 249 | apply (drule zless_trans, assumption) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 250 | apply (auto simp add: zero_lt_zmagnitude) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 251 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 252 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 253 | lemma zmult_zless_mono1: "[| i$<j; #0 $< k |] ==> i$*k $< j$*k" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 254 | apply (drule zmult_zless_mono2) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 255 | apply (simp_all add: zmult_commute) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 256 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 257 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 258 | (* < monotonicity, BOTH arguments*) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 259 | lemma zmult_zless_mono: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 260 | "[| i $< j; k $< l; #0 $< j; #0 $< k |] ==> i$*k $< j$*l" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 261 | apply (erule zmult_zless_mono1 [THEN zless_trans]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 262 | apply assumption | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 263 | apply (erule zmult_zless_mono2) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 264 | apply assumption | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 265 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 266 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 267 | lemma zmult_zless_mono1_neg: "[| i $< j; k $< #0 |] ==> j$*k $< i$*k" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 268 | apply (rule zminus_zless_zminus [THEN iffD1]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 269 | apply (simp del: zmult_zminus_right | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 270 | add: zmult_zminus_right [symmetric] zmult_zless_mono1 zless_zminus) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 271 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 272 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 273 | lemma zmult_zless_mono2_neg: "[| i $< j; k $< #0 |] ==> k$*j $< k$*i" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 274 | apply (rule zminus_zless_zminus [THEN iffD1]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 275 | apply (simp del: zmult_zminus | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 276 | add: zmult_zminus [symmetric] zmult_zless_mono2 zless_zminus) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 277 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 278 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 279 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 280 | (** Products of zeroes **) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 281 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 282 | lemma zmult_eq_lemma: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 283 | "[| m \<in> int; n \<in> int |] ==> (m = #0 | n = #0) <-> (m$*n = #0)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 284 | apply (case_tac "m $< #0") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 285 | apply (auto simp add: not_zless_iff_zle zle_def neq_iff_zless) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 286 | apply (force dest: zmult_zless_mono1_neg zmult_zless_mono1)+ | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 287 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 288 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 289 | lemma zmult_eq_0_iff [iff]: "(m$*n = #0) <-> (intify(m) = #0 | intify(n) = #0)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 290 | apply (simp add: zmult_eq_lemma) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 291 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 292 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 293 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 294 | (** Cancellation laws for k*m < k*n and m*k < n*k, also for <= and =, | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 295 | but not (yet?) for k*m < n*k. **) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 296 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 297 | lemma zmult_zless_lemma: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 298 | "[| k \<in> int; m \<in> int; n \<in> int |] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 299 | ==> (m$*k $< n$*k) <-> ((#0 $< k & m$<n) | (k $< #0 & n$<m))" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 300 | apply (case_tac "k = #0") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 301 | apply (auto simp add: neq_iff_zless zmult_zless_mono1 zmult_zless_mono1_neg) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 302 | apply (auto simp add: not_zless_iff_zle | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 303 | not_zle_iff_zless [THEN iff_sym, of "m$*k"] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 304 | not_zle_iff_zless [THEN iff_sym, of m]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 305 | apply (auto elim: notE | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 306 | simp add: zless_imp_zle zmult_zle_mono1 zmult_zle_mono1_neg) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 307 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 308 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 309 | lemma zmult_zless_cancel2: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 310 | "(m$*k $< n$*k) <-> ((#0 $< k & m$<n) | (k $< #0 & n$<m))" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 311 | apply (cut_tac k = "intify (k)" and m = "intify (m)" and n = "intify (n)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 312 | in zmult_zless_lemma) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 313 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 314 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 315 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 316 | lemma zmult_zless_cancel1: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 317 | "(k$*m $< k$*n) <-> ((#0 $< k & m$<n) | (k $< #0 & n$<m))" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 318 | by (simp add: zmult_commute [of k] zmult_zless_cancel2) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 319 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 320 | lemma zmult_zle_cancel2: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 321 | "(m$*k $<= n$*k) <-> ((#0 $< k --> m$<=n) & (k $< #0 --> n$<=m))" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 322 | by (auto simp add: not_zless_iff_zle [THEN iff_sym] zmult_zless_cancel2) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 323 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 324 | lemma zmult_zle_cancel1: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 325 | "(k$*m $<= k$*n) <-> ((#0 $< k --> m$<=n) & (k $< #0 --> n$<=m))" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 326 | by (auto simp add: not_zless_iff_zle [THEN iff_sym] zmult_zless_cancel1) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 327 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 328 | lemma int_eq_iff_zle: "[| m \<in> int; n \<in> int |] ==> m=n <-> (m $<= n & n $<= m)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 329 | apply (blast intro: zle_refl zle_anti_sym) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 330 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 331 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 332 | lemma zmult_cancel2_lemma: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 333 | "[| k \<in> int; m \<in> int; n \<in> int |] ==> (m$*k = n$*k) <-> (k=#0 | m=n)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 334 | apply (simp add: int_eq_iff_zle [of "m$*k"] int_eq_iff_zle [of m]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 335 | apply (auto simp add: zmult_zle_cancel2 neq_iff_zless) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 336 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 337 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 338 | lemma zmult_cancel2 [simp]: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 339 | "(m$*k = n$*k) <-> (intify(k) = #0 | intify(m) = intify(n))" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 340 | apply (rule iff_trans) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 341 | apply (rule_tac [2] zmult_cancel2_lemma) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 342 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 343 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 344 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 345 | lemma zmult_cancel1 [simp]: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 346 | "(k$*m = k$*n) <-> (intify(k) = #0 | intify(m) = intify(n))" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 347 | by (simp add: zmult_commute [of k] zmult_cancel2) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 348 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 349 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 350 | subsection{* Uniqueness and monotonicity of quotients and remainders *}
 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 351 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 352 | lemma unique_quotient_lemma: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 353 | "[| b$*q' $+ r' $<= b$*q $+ r; #0 $<= r'; #0 $< b; r $< b |] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 354 | ==> q' $<= q" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 355 | apply (subgoal_tac "r' $+ b $* (q'$-q) $<= r") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 356 | prefer 2 apply (simp add: zdiff_zmult_distrib2 zadd_ac zcompare_rls) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 357 | apply (subgoal_tac "#0 $< b $* (#1 $+ q $- q') ") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 358 | prefer 2 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 359 | apply (erule zle_zless_trans) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 360 | apply (simp add: zdiff_zmult_distrib2 zadd_zmult_distrib2 zadd_ac zcompare_rls) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 361 | apply (erule zle_zless_trans) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 362 | apply (simp add: ); | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 363 | apply (subgoal_tac "b $* q' $< b $* (#1 $+ q)") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 364 | prefer 2 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 365 | apply (simp add: zdiff_zmult_distrib2 zadd_zmult_distrib2 zadd_ac zcompare_rls) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 366 | apply (auto elim: zless_asym | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 367 | simp add: zmult_zless_cancel1 zless_add1_iff_zle zadd_ac zcompare_rls) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 368 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 369 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 370 | lemma unique_quotient_lemma_neg: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 371 | "[| b$*q' $+ r' $<= b$*q $+ r; r $<= #0; b $< #0; b $< r' |] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 372 | ==> q $<= q'" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 373 | apply (rule_tac b = "$-b" and r = "$-r'" and r' = "$-r" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 374 | in unique_quotient_lemma) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 375 | apply (auto simp del: zminus_zadd_distrib | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 376 | simp add: zminus_zadd_distrib [symmetric] zle_zminus zless_zminus) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 377 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 378 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 379 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 380 | lemma unique_quotient: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 381 | "[| quorem (<a,b>, <q,r>); quorem (<a,b>, <q',r'>); b \<in> int; b ~= #0; | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 382 | q \<in> int; q' \<in> int |] ==> q = q'" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 383 | apply (simp add: split_ifs quorem_def neq_iff_zless) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 384 | apply safe | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 385 | apply simp_all | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 386 | apply (blast intro: zle_anti_sym | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 387 | dest: zle_eq_refl [THEN unique_quotient_lemma] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 388 | zle_eq_refl [THEN unique_quotient_lemma_neg] sym)+ | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 389 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 390 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 391 | lemma unique_remainder: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 392 | "[| quorem (<a,b>, <q,r>); quorem (<a,b>, <q',r'>); b \<in> int; b ~= #0; | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 393 | q \<in> int; q' \<in> int; | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 394 | r \<in> int; r' \<in> int |] ==> r = r'" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 395 | apply (subgoal_tac "q = q'") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 396 | prefer 2 apply (blast intro: unique_quotient) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 397 | apply (simp add: quorem_def) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 398 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 399 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 400 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 401 | subsection{*Correctness of posDivAlg, 
 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 402 |            the Division Algorithm for @{text "a\<ge>0"} and @{text "b>0"} *}
 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 403 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 404 | lemma adjust_eq [simp]: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 405 | "adjust(b, <q,r>) = (let diff = r$-b in | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 406 | if #0 $<= diff then <#2$*q $+ #1,diff> | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 407 | else <#2$*q,r>)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 408 | by (simp add: Let_def adjust_def) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 409 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 410 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 411 | lemma posDivAlg_termination: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 412 | "[| #0 $< b; ~ a $< b |] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 413 | ==> nat_of(a $- #2 $\<times> b $+ #1) < nat_of(a $- b $+ #1)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 414 | apply (simp (no_asm) add: zless_nat_conj) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 415 | apply (simp add: not_zless_iff_zle zless_add1_iff_zle zcompare_rls) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 416 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 417 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 418 | lemmas posDivAlg_unfold = def_wfrec [OF posDivAlg_def wf_measure] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 419 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 420 | lemma posDivAlg_eqn: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 421 | "[| #0 $< b; a \<in> int; b \<in> int |] ==> | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 422 | posDivAlg(<a,b>) = | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 423 | (if a$<b then <#0,a> else adjust(b, posDivAlg (<a, #2$*b>)))" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 424 | apply (rule posDivAlg_unfold [THEN trans]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 425 | apply (simp add: vimage_iff not_zless_iff_zle [THEN iff_sym]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 426 | apply (blast intro: posDivAlg_termination) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 427 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 428 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 429 | lemma posDivAlg_induct_lemma [rule_format]: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 430 | assumes prem: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 431 | "!!a b. [| a \<in> int; b \<in> int; | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 432 | ~ (a $< b | b $<= #0) --> P(<a, #2 $* b>) |] ==> P(<a,b>)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 433 | shows "<u,v> \<in> int*int --> P(<u,v>)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 434 | apply (rule_tac a = "<u,v>" in wf_induct) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 435 | apply (rule_tac A = "int*int" and f = "%<a,b>.nat_of (a $- b $+ #1)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 436 | in wf_measure) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 437 | apply clarify | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 438 | apply (rule prem) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 439 | apply (drule_tac [3] x = "<xa, #2 $\<times> y>" in spec) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 440 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 441 | apply (simp add: not_zle_iff_zless posDivAlg_termination) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 442 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 443 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 444 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 445 | lemma posDivAlg_induct [consumes 2]: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 446 | assumes u_int: "u \<in> int" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 447 | and v_int: "v \<in> int" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 448 | and ih: "!!a b. [| a \<in> int; b \<in> int; | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 449 | ~ (a $< b | b $<= #0) --> P(a, #2 $* b) |] ==> P(a,b)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 450 | shows "P(u,v)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 451 | apply (subgoal_tac "(%<x,y>. P (x,y)) (<u,v>)") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 452 | apply simp | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 453 | apply (rule posDivAlg_induct_lemma) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 454 | apply (simp (no_asm_use)) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 455 | apply (rule ih) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 456 | apply (auto simp add: u_int v_int) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 457 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 458 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 459 | (*FIXME: use intify in integ_of so that we always have integ_of w \<in> int. | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 460 | then this rewrite can work for ALL constants!!*) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 461 | lemma intify_eq_0_iff_zle: "intify(m) = #0 <-> (m $<= #0 & #0 $<= m)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 462 | apply (simp (no_asm) add: int_eq_iff_zle) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 463 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 464 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 465 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 466 | subsection{* Some convenient biconditionals for products of signs *}
 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 467 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 468 | lemma zmult_pos: "[| #0 $< i; #0 $< j |] ==> #0 $< i $* j" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 469 | apply (drule zmult_zless_mono1) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 470 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 471 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 472 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 473 | lemma zmult_neg: "[| i $< #0; j $< #0 |] ==> #0 $< i $* j" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 474 | apply (drule zmult_zless_mono1_neg) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 475 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 476 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 477 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 478 | lemma zmult_pos_neg: "[| #0 $< i; j $< #0 |] ==> i $* j $< #0" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 479 | apply (drule zmult_zless_mono1_neg) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 480 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 481 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 482 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 483 | (** Inequality reasoning **) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 484 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 485 | lemma int_0_less_lemma: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 486 | "[| x \<in> int; y \<in> int |] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 487 | ==> (#0 $< x $* y) <-> (#0 $< x & #0 $< y | x $< #0 & y $< #0)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 488 | apply (auto simp add: zle_def not_zless_iff_zle zmult_pos zmult_neg) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 489 | apply (rule ccontr) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 490 | apply (rule_tac [2] ccontr) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 491 | apply (auto simp add: zle_def not_zless_iff_zle) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 492 | apply (erule_tac P = "#0$< x$* y" in rev_mp) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 493 | apply (erule_tac [2] P = "#0$< x$* y" in rev_mp) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 494 | apply (drule zmult_pos_neg, assumption) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 495 | prefer 2 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 496 | apply (drule zmult_pos_neg, assumption) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 497 | apply (auto dest: zless_not_sym simp add: zmult_commute) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 498 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 499 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 500 | lemma int_0_less_mult_iff: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 501 | "(#0 $< x $* y) <-> (#0 $< x & #0 $< y | x $< #0 & y $< #0)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 502 | apply (cut_tac x = "intify (x)" and y = "intify (y)" in int_0_less_lemma) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 503 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 504 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 505 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 506 | lemma int_0_le_lemma: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 507 | "[| x \<in> int; y \<in> int |] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 508 | ==> (#0 $<= x $* y) <-> (#0 $<= x & #0 $<= y | x $<= #0 & y $<= #0)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 509 | by (auto simp add: zle_def not_zless_iff_zle int_0_less_mult_iff) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 510 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 511 | lemma int_0_le_mult_iff: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 512 | "(#0 $<= x $* y) <-> ((#0 $<= x & #0 $<= y) | (x $<= #0 & y $<= #0))" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 513 | apply (cut_tac x = "intify (x)" and y = "intify (y)" in int_0_le_lemma) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 514 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 515 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 516 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 517 | lemma zmult_less_0_iff: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 518 | "(x $* y $< #0) <-> (#0 $< x & y $< #0 | x $< #0 & #0 $< y)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 519 | apply (auto simp add: int_0_le_mult_iff not_zle_iff_zless [THEN iff_sym]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 520 | apply (auto dest: zless_not_sym simp add: not_zle_iff_zless) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 521 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 522 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 523 | lemma zmult_le_0_iff: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 524 | "(x $* y $<= #0) <-> (#0 $<= x & y $<= #0 | x $<= #0 & #0 $<= y)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 525 | by (auto dest: zless_not_sym | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 526 | simp add: int_0_less_mult_iff not_zless_iff_zle [THEN iff_sym]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 527 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 528 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 529 | (*Typechecking for posDivAlg*) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 530 | lemma posDivAlg_type [rule_format]: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 531 | "[| a \<in> int; b \<in> int |] ==> posDivAlg(<a,b>) \<in> int * int" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 532 | apply (rule_tac u = "a" and v = "b" in posDivAlg_induct) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 533 | apply assumption+ | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 534 | apply (case_tac "#0 $< ba") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 535 | apply (simp add: posDivAlg_eqn adjust_def integ_of_type | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 536 | split add: split_if_asm) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 537 | apply clarify | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 538 | apply (simp add: int_0_less_mult_iff not_zle_iff_zless) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 539 | apply (simp add: not_zless_iff_zle) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 540 | apply (subst posDivAlg_unfold) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 541 | apply simp | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 542 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 543 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 544 | (*Correctness of posDivAlg: it computes quotients correctly*) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 545 | lemma posDivAlg_correct [rule_format]: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 546 | "[| a \<in> int; b \<in> int |] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 547 | ==> #0 $<= a --> #0 $< b --> quorem (<a,b>, posDivAlg(<a,b>))" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 548 | apply (rule_tac u = "a" and v = "b" in posDivAlg_induct) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 549 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 550 | apply (simp_all add: quorem_def) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 551 |    txt{*base case: a<b*}
 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 552 | apply (simp add: posDivAlg_eqn) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 553 | apply (simp add: not_zless_iff_zle [THEN iff_sym]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 554 | apply (simp add: int_0_less_mult_iff) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 555 | txt{*main argument*}
 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 556 | apply (subst posDivAlg_eqn) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 557 | apply (simp_all (no_asm_simp)) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 558 | apply (erule splitE) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 559 | apply (rule posDivAlg_type) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 560 | apply (simp_all add: int_0_less_mult_iff) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 561 | apply (auto simp add: zadd_zmult_distrib2 Let_def) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 562 | txt{*now just linear arithmetic*}
 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 563 | apply (simp add: not_zle_iff_zless zdiff_zless_iff) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 564 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 565 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 566 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 567 | subsection{*Correctness of negDivAlg, the division algorithm for a<0 and b>0*}
 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 568 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 569 | lemma negDivAlg_termination: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 570 | "[| #0 $< b; a $+ b $< #0 |] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 571 | ==> nat_of($- a $- #2 $* b) < nat_of($- a $- b)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 572 | apply (simp (no_asm) add: zless_nat_conj) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 573 | apply (simp add: zcompare_rls not_zle_iff_zless zless_zdiff_iff [THEN iff_sym] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 574 | zless_zminus) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 575 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 576 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 577 | lemmas negDivAlg_unfold = def_wfrec [OF negDivAlg_def wf_measure] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 578 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 579 | lemma negDivAlg_eqn: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 580 | "[| #0 $< b; a : int; b : int |] ==> | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 581 | negDivAlg(<a,b>) = | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 582 | (if #0 $<= a$+b then <#-1,a$+b> | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 583 | else adjust(b, negDivAlg (<a, #2$*b>)))" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 584 | apply (rule negDivAlg_unfold [THEN trans]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 585 | apply (simp (no_asm_simp) add: vimage_iff not_zless_iff_zle [THEN iff_sym]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 586 | apply (blast intro: negDivAlg_termination) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 587 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 588 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 589 | lemma negDivAlg_induct_lemma [rule_format]: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 590 | assumes prem: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 591 | "!!a b. [| a \<in> int; b \<in> int; | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 592 | ~ (#0 $<= a $+ b | b $<= #0) --> P(<a, #2 $* b>) |] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 593 | ==> P(<a,b>)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 594 | shows "<u,v> \<in> int*int --> P(<u,v>)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 595 | apply (rule_tac a = "<u,v>" in wf_induct) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 596 | apply (rule_tac A = "int*int" and f = "%<a,b>.nat_of ($- a $- b)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 597 | in wf_measure) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 598 | apply clarify | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 599 | apply (rule prem) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 600 | apply (drule_tac [3] x = "<xa, #2 $\<times> y>" in spec) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 601 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 602 | apply (simp add: not_zle_iff_zless negDivAlg_termination) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 603 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 604 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 605 | lemma negDivAlg_induct [consumes 2]: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 606 | assumes u_int: "u \<in> int" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 607 | and v_int: "v \<in> int" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 608 | and ih: "!!a b. [| a \<in> int; b \<in> int; | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 609 | ~ (#0 $<= a $+ b | b $<= #0) --> P(a, #2 $* b) |] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 610 | ==> P(a,b)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 611 | shows "P(u,v)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 612 | apply (subgoal_tac " (%<x,y>. P (x,y)) (<u,v>)") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 613 | apply simp | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 614 | apply (rule negDivAlg_induct_lemma) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 615 | apply (simp (no_asm_use)) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 616 | apply (rule ih) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 617 | apply (auto simp add: u_int v_int) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 618 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 619 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 620 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 621 | (*Typechecking for negDivAlg*) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 622 | lemma negDivAlg_type: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 623 | "[| a \<in> int; b \<in> int |] ==> negDivAlg(<a,b>) \<in> int * int" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 624 | apply (rule_tac u = "a" and v = "b" in negDivAlg_induct) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 625 | apply assumption+ | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 626 | apply (case_tac "#0 $< ba") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 627 | apply (simp add: negDivAlg_eqn adjust_def integ_of_type | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 628 | split add: split_if_asm) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 629 | apply clarify | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 630 | apply (simp add: int_0_less_mult_iff not_zle_iff_zless) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 631 | apply (simp add: not_zless_iff_zle) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 632 | apply (subst negDivAlg_unfold) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 633 | apply simp | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 634 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 635 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 636 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 637 | (*Correctness of negDivAlg: it computes quotients correctly | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 638 | It doesn't work if a=0 because the 0/b=0 rather than -1*) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 639 | lemma negDivAlg_correct [rule_format]: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 640 | "[| a \<in> int; b \<in> int |] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 641 | ==> a $< #0 --> #0 $< b --> quorem (<a,b>, negDivAlg(<a,b>))" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 642 | apply (rule_tac u = "a" and v = "b" in negDivAlg_induct) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 643 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 644 | apply (simp_all add: quorem_def) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 645 |    txt{*base case: @{term "0$<=a$+b"}*}
 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 646 | apply (simp add: negDivAlg_eqn) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 647 | apply (simp add: not_zless_iff_zle [THEN iff_sym]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 648 | apply (simp add: int_0_less_mult_iff) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 649 | txt{*main argument*}
 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 650 | apply (subst negDivAlg_eqn) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 651 | apply (simp_all (no_asm_simp)) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 652 | apply (erule splitE) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 653 | apply (rule negDivAlg_type) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 654 | apply (simp_all add: int_0_less_mult_iff) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 655 | apply (auto simp add: zadd_zmult_distrib2 Let_def) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 656 | txt{*now just linear arithmetic*}
 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 657 | apply (simp add: not_zle_iff_zless zdiff_zless_iff) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 658 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 659 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 660 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 661 | subsection{* Existence shown by proving the division algorithm to be correct *}
 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 662 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 663 | (*the case a=0*) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 664 | lemma quorem_0: "[|b \<noteq> #0; b \<in> int|] ==> quorem (<#0,b>, <#0,#0>)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 665 | by (force simp add: quorem_def neq_iff_zless) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 666 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 667 | lemma posDivAlg_zero_divisor: "posDivAlg(<a,#0>) = <#0,a>" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 668 | apply (subst posDivAlg_unfold) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 669 | apply simp | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 670 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 671 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 672 | lemma posDivAlg_0 [simp]: "posDivAlg (<#0,b>) = <#0,#0>" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 673 | apply (subst posDivAlg_unfold) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 674 | apply (simp add: not_zle_iff_zless) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 675 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 676 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 677 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 678 | (*Needed below. Actually it's an equivalence.*) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 679 | lemma linear_arith_lemma: "~ (#0 $<= #-1 $+ b) ==> (b $<= #0)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 680 | apply (simp add: not_zle_iff_zless) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 681 | apply (drule zminus_zless_zminus [THEN iffD2]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 682 | apply (simp add: zadd_commute zless_add1_iff_zle zle_zminus) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 683 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 684 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 685 | lemma negDivAlg_minus1 [simp]: "negDivAlg (<#-1,b>) = <#-1, b$-#1>" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 686 | apply (subst negDivAlg_unfold) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 687 | apply (simp add: linear_arith_lemma integ_of_type vimage_iff) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 688 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 689 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 690 | lemma negateSnd_eq [simp]: "negateSnd (<q,r>) = <q, $-r>" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 691 | apply (unfold negateSnd_def) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 692 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 693 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 694 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 695 | lemma negateSnd_type: "qr \<in> int * int ==> negateSnd (qr) \<in> int * int" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 696 | apply (unfold negateSnd_def) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 697 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 698 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 699 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 700 | lemma quorem_neg: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 701 | "[|quorem (<$-a,$-b>, qr); a \<in> int; b \<in> int; qr \<in> int * int|] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 702 | ==> quorem (<a,b>, negateSnd(qr))" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 703 | apply clarify | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 704 | apply (auto elim: zless_asym simp add: quorem_def zless_zminus) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 705 | txt{*linear arithmetic from here on*}
 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 706 | apply (simp_all add: zminus_equation [of a] zminus_zless) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 707 | apply (cut_tac [2] z = "b" and w = "#0" in zless_linear) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 708 | apply (cut_tac [1] z = "b" and w = "#0" in zless_linear) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 709 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 710 | apply (blast dest: zle_zless_trans)+ | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 711 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 712 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 713 | lemma divAlg_correct: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 714 | "[|b \<noteq> #0; a \<in> int; b \<in> int|] ==> quorem (<a,b>, divAlg(<a,b>))" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 715 | apply (auto simp add: quorem_0 divAlg_def) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 716 | apply (safe intro!: quorem_neg posDivAlg_correct negDivAlg_correct | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 717 | posDivAlg_type negDivAlg_type) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 718 | apply (auto simp add: quorem_def neq_iff_zless) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 719 | txt{*linear arithmetic from here on*}
 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 720 | apply (auto simp add: zle_def) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 721 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 722 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 723 | lemma divAlg_type: "[|a \<in> int; b \<in> int|] ==> divAlg(<a,b>) \<in> int * int" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 724 | apply (auto simp add: divAlg_def) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 725 | apply (auto simp add: posDivAlg_type negDivAlg_type negateSnd_type) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 726 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 727 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 728 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 729 | (** intify cancellation **) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 730 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 731 | lemma zdiv_intify1 [simp]: "intify(x) zdiv y = x zdiv y" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 732 | apply (simp (no_asm) add: zdiv_def) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 733 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 734 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 735 | lemma zdiv_intify2 [simp]: "x zdiv intify(y) = x zdiv y" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 736 | apply (simp (no_asm) add: zdiv_def) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 737 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 738 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 739 | lemma zdiv_type [iff,TC]: "z zdiv w \<in> int" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 740 | apply (unfold zdiv_def) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 741 | apply (blast intro: fst_type divAlg_type) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 742 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 743 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 744 | lemma zmod_intify1 [simp]: "intify(x) zmod y = x zmod y" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 745 | apply (simp (no_asm) add: zmod_def) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 746 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 747 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 748 | lemma zmod_intify2 [simp]: "x zmod intify(y) = x zmod y" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 749 | apply (simp (no_asm) add: zmod_def) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 750 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 751 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 752 | lemma zmod_type [iff,TC]: "z zmod w \<in> int" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 753 | apply (unfold zmod_def) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 754 | apply (rule snd_type) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 755 | apply (blast intro: divAlg_type) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 756 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 757 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 758 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 759 | (** Arbitrary definitions for division by zero. Useful to simplify | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 760 | certain equations **) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 761 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 762 | lemma DIVISION_BY_ZERO_ZDIV: "a zdiv #0 = #0" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 763 | apply (simp (no_asm) add: zdiv_def divAlg_def posDivAlg_zero_divisor) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 764 | done (*NOT for adding to default simpset*) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 765 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 766 | lemma DIVISION_BY_ZERO_ZMOD: "a zmod #0 = intify(a)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 767 | apply (simp (no_asm) add: zmod_def divAlg_def posDivAlg_zero_divisor) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 768 | done (*NOT for adding to default simpset*) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 769 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 770 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 771 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 772 | (** Basic laws about division and remainder **) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 773 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 774 | lemma raw_zmod_zdiv_equality: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 775 | "[| a \<in> int; b \<in> int |] ==> a = b $* (a zdiv b) $+ (a zmod b)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 776 | apply (case_tac "b = #0") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 777 | apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 778 | apply (cut_tac a = "a" and b = "b" in divAlg_correct) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 779 | apply (auto simp add: quorem_def zdiv_def zmod_def split_def) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 780 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 781 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 782 | lemma zmod_zdiv_equality: "intify(a) = b $* (a zdiv b) $+ (a zmod b)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 783 | apply (rule trans) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 784 | apply (rule_tac b = "intify (b)" in raw_zmod_zdiv_equality) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 785 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 786 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 787 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 788 | lemma pos_mod: "#0 $< b ==> #0 $<= a zmod b & a zmod b $< b" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 789 | apply (cut_tac a = "intify (a)" and b = "intify (b)" in divAlg_correct) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 790 | apply (auto simp add: intify_eq_0_iff_zle quorem_def zmod_def split_def) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 791 | apply (blast dest: zle_zless_trans)+ | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 792 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 793 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 794 | lemmas pos_mod_sign = pos_mod [THEN conjunct1, standard] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 795 | and pos_mod_bound = pos_mod [THEN conjunct2, standard] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 796 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 797 | lemma neg_mod: "b $< #0 ==> a zmod b $<= #0 & b $< a zmod b" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 798 | apply (cut_tac a = "intify (a)" and b = "intify (b)" in divAlg_correct) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 799 | apply (auto simp add: intify_eq_0_iff_zle quorem_def zmod_def split_def) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 800 | apply (blast dest: zle_zless_trans) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 801 | apply (blast dest: zless_trans)+ | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 802 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 803 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 804 | lemmas neg_mod_sign = neg_mod [THEN conjunct1, standard] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 805 | and neg_mod_bound = neg_mod [THEN conjunct2, standard] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 806 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 807 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 808 | (** proving general properties of zdiv and zmod **) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 809 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 810 | lemma quorem_div_mod: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 811 | "[|b \<noteq> #0; a \<in> int; b \<in> int |] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 812 | ==> quorem (<a,b>, <a zdiv b, a zmod b>)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 813 | apply (cut_tac a = "a" and b = "b" in zmod_zdiv_equality) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 814 | apply (auto simp add: quorem_def neq_iff_zless pos_mod_sign pos_mod_bound | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 815 | neg_mod_sign neg_mod_bound) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 816 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 817 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 818 | (*Surely quorem(<a,b>,<q,r>) implies a \<in> int, but it doesn't matter*) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 819 | lemma quorem_div: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 820 | "[| quorem(<a,b>,<q,r>); b \<noteq> #0; a \<in> int; b \<in> int; q \<in> int |] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 821 | ==> a zdiv b = q" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 822 | by (blast intro: quorem_div_mod [THEN unique_quotient]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 823 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 824 | lemma quorem_mod: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 825 | "[| quorem(<a,b>,<q,r>); b \<noteq> #0; a \<in> int; b \<in> int; q \<in> int; r \<in> int |] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 826 | ==> a zmod b = r" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 827 | by (blast intro: quorem_div_mod [THEN unique_remainder]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 828 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 829 | lemma zdiv_pos_pos_trivial_raw: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 830 | "[| a \<in> int; b \<in> int; #0 $<= a; a $< b |] ==> a zdiv b = #0" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 831 | apply (rule quorem_div) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 832 | apply (auto simp add: quorem_def) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 833 | (*linear arithmetic*) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 834 | apply (blast dest: zle_zless_trans)+ | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 835 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 836 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 837 | lemma zdiv_pos_pos_trivial: "[| #0 $<= a; a $< b |] ==> a zdiv b = #0" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 838 | apply (cut_tac a = "intify (a)" and b = "intify (b)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 839 | in zdiv_pos_pos_trivial_raw) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 840 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 841 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 842 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 843 | lemma zdiv_neg_neg_trivial_raw: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 844 | "[| a \<in> int; b \<in> int; a $<= #0; b $< a |] ==> a zdiv b = #0" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 845 | apply (rule_tac r = "a" in quorem_div) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 846 | apply (auto simp add: quorem_def) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 847 | (*linear arithmetic*) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 848 | apply (blast dest: zle_zless_trans zless_trans)+ | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 849 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 850 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 851 | lemma zdiv_neg_neg_trivial: "[| a $<= #0; b $< a |] ==> a zdiv b = #0" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 852 | apply (cut_tac a = "intify (a)" and b = "intify (b)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 853 | in zdiv_neg_neg_trivial_raw) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 854 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 855 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 856 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 857 | lemma zadd_le_0_lemma: "[| a$+b $<= #0; #0 $< a; #0 $< b |] ==> False" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 858 | apply (drule_tac z' = "#0" and z = "b" in zadd_zless_mono) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 859 | apply (auto simp add: zle_def) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 860 | apply (blast dest: zless_trans) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 861 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 862 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 863 | lemma zdiv_pos_neg_trivial_raw: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 864 | "[| a \<in> int; b \<in> int; #0 $< a; a$+b $<= #0 |] ==> a zdiv b = #-1" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 865 | apply (rule_tac r = "a $+ b" in quorem_div) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 866 | apply (auto simp add: quorem_def) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 867 | (*linear arithmetic*) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 868 | apply (blast dest: zadd_le_0_lemma zle_zless_trans)+ | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 869 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 870 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 871 | lemma zdiv_pos_neg_trivial: "[| #0 $< a; a$+b $<= #0 |] ==> a zdiv b = #-1" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 872 | apply (cut_tac a = "intify (a)" and b = "intify (b)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 873 | in zdiv_pos_neg_trivial_raw) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 874 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 875 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 876 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 877 | (*There is no zdiv_neg_pos_trivial because #0 zdiv b = #0 would supersede it*) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 878 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 879 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 880 | lemma zmod_pos_pos_trivial_raw: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 881 | "[| a \<in> int; b \<in> int; #0 $<= a; a $< b |] ==> a zmod b = a" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 882 | apply (rule_tac q = "#0" in quorem_mod) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 883 | apply (auto simp add: quorem_def) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 884 | (*linear arithmetic*) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 885 | apply (blast dest: zle_zless_trans)+ | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 886 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 887 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 888 | lemma zmod_pos_pos_trivial: "[| #0 $<= a; a $< b |] ==> a zmod b = intify(a)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 889 | apply (cut_tac a = "intify (a)" and b = "intify (b)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 890 | in zmod_pos_pos_trivial_raw) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 891 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 892 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 893 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 894 | lemma zmod_neg_neg_trivial_raw: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 895 | "[| a \<in> int; b \<in> int; a $<= #0; b $< a |] ==> a zmod b = a" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 896 | apply (rule_tac q = "#0" in quorem_mod) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 897 | apply (auto simp add: quorem_def) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 898 | (*linear arithmetic*) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 899 | apply (blast dest: zle_zless_trans zless_trans)+ | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 900 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 901 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 902 | lemma zmod_neg_neg_trivial: "[| a $<= #0; b $< a |] ==> a zmod b = intify(a)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 903 | apply (cut_tac a = "intify (a)" and b = "intify (b)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 904 | in zmod_neg_neg_trivial_raw) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 905 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 906 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 907 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 908 | lemma zmod_pos_neg_trivial_raw: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 909 | "[| a \<in> int; b \<in> int; #0 $< a; a$+b $<= #0 |] ==> a zmod b = a$+b" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 910 | apply (rule_tac q = "#-1" in quorem_mod) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 911 | apply (auto simp add: quorem_def) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 912 | (*linear arithmetic*) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 913 | apply (blast dest: zadd_le_0_lemma zle_zless_trans)+ | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 914 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 915 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 916 | lemma zmod_pos_neg_trivial: "[| #0 $< a; a$+b $<= #0 |] ==> a zmod b = a$+b" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 917 | apply (cut_tac a = "intify (a)" and b = "intify (b)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 918 | in zmod_pos_neg_trivial_raw) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 919 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 920 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 921 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 922 | (*There is no zmod_neg_pos_trivial...*) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 923 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 924 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 925 | (*Simpler laws such as -a zdiv b = -(a zdiv b) FAIL*) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 926 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 927 | lemma zdiv_zminus_zminus_raw: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 928 | "[|a \<in> int; b \<in> int|] ==> ($-a) zdiv ($-b) = a zdiv b" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 929 | apply (case_tac "b = #0") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 930 | apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 931 | apply (subst quorem_div_mod [THEN quorem_neg, simplified, THEN quorem_div]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 932 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 933 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 934 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 935 | lemma zdiv_zminus_zminus [simp]: "($-a) zdiv ($-b) = a zdiv b" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 936 | apply (cut_tac a = "intify (a)" and b = "intify (b)" in zdiv_zminus_zminus_raw) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 937 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 938 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 939 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 940 | (*Simpler laws such as -a zmod b = -(a zmod b) FAIL*) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 941 | lemma zmod_zminus_zminus_raw: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 942 | "[|a \<in> int; b \<in> int|] ==> ($-a) zmod ($-b) = $- (a zmod b)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 943 | apply (case_tac "b = #0") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 944 | apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 945 | apply (subst quorem_div_mod [THEN quorem_neg, simplified, THEN quorem_mod]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 946 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 947 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 948 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 949 | lemma zmod_zminus_zminus [simp]: "($-a) zmod ($-b) = $- (a zmod b)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 950 | apply (cut_tac a = "intify (a)" and b = "intify (b)" in zmod_zminus_zminus_raw) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 951 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 952 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 953 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 954 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 955 | subsection{* division of a number by itself *}
 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 956 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 957 | lemma self_quotient_aux1: "[| #0 $< a; a = r $+ a$*q; r $< a |] ==> #1 $<= q" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 958 | apply (subgoal_tac "#0 $< a$*q") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 959 | apply (cut_tac w = "#0" and z = "q" in add1_zle_iff) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 960 | apply (simp add: int_0_less_mult_iff) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 961 | apply (blast dest: zless_trans) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 962 | (*linear arithmetic...*) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 963 | apply (drule_tac t = "%x. x $- r" in subst_context) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 964 | apply (drule sym) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 965 | apply (simp add: zcompare_rls) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 966 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 967 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 968 | lemma self_quotient_aux2: "[| #0 $< a; a = r $+ a$*q; #0 $<= r |] ==> q $<= #1" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 969 | apply (subgoal_tac "#0 $<= a$* (#1$-q)") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 970 | apply (simp add: int_0_le_mult_iff zcompare_rls) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 971 | apply (blast dest: zle_zless_trans) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 972 | apply (simp add: zdiff_zmult_distrib2) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 973 | apply (drule_tac t = "%x. x $- a $* q" in subst_context) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 974 | apply (simp add: zcompare_rls) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 975 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 976 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 977 | lemma self_quotient: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 978 | "[| quorem(<a,a>,<q,r>); a \<in> int; q \<in> int; a \<noteq> #0|] ==> q = #1" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 979 | apply (simp add: split_ifs quorem_def neq_iff_zless) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 980 | apply (rule zle_anti_sym) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 981 | apply safe | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 982 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 983 | prefer 4 apply (blast dest: zless_trans) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 984 | apply (blast dest: zless_trans) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 985 | apply (rule_tac [3] a = "$-a" and r = "$-r" in self_quotient_aux1) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 986 | apply (rule_tac a = "$-a" and r = "$-r" in self_quotient_aux2) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 987 | apply (rule_tac [6] zminus_equation [THEN iffD1]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 988 | apply (rule_tac [2] zminus_equation [THEN iffD1]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 989 | apply (force intro: self_quotient_aux1 self_quotient_aux2 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 990 | simp add: zadd_commute zmult_zminus)+ | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 991 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 992 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 993 | lemma self_remainder: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 994 | "[|quorem(<a,a>,<q,r>); a \<in> int; q \<in> int; r \<in> int; a \<noteq> #0|] ==> r = #0" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 995 | apply (frule self_quotient) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 996 | apply (auto simp add: quorem_def) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 997 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 998 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 999 | lemma zdiv_self_raw: "[|a \<noteq> #0; a \<in> int|] ==> a zdiv a = #1" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1000 | apply (blast intro: quorem_div_mod [THEN self_quotient]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1001 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1002 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1003 | lemma zdiv_self [simp]: "intify(a) \<noteq> #0 ==> a zdiv a = #1" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1004 | apply (drule zdiv_self_raw) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1005 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1006 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1007 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1008 | (*Here we have 0 zmod 0 = 0, also assumed by Knuth (who puts m zmod 0 = 0) *) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1009 | lemma zmod_self_raw: "a \<in> int ==> a zmod a = #0" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1010 | apply (case_tac "a = #0") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1011 | apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1012 | apply (blast intro: quorem_div_mod [THEN self_remainder]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1013 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1014 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1015 | lemma zmod_self [simp]: "a zmod a = #0" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1016 | apply (cut_tac a = "intify (a)" in zmod_self_raw) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1017 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1018 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1019 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1020 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1021 | subsection{* Computation of division and remainder *}
 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1022 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1023 | lemma zdiv_zero [simp]: "#0 zdiv b = #0" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1024 | apply (simp (no_asm) add: zdiv_def divAlg_def) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1025 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1026 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1027 | lemma zdiv_eq_minus1: "#0 $< b ==> #-1 zdiv b = #-1" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1028 | apply (simp (no_asm_simp) add: zdiv_def divAlg_def) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1029 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1030 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1031 | lemma zmod_zero [simp]: "#0 zmod b = #0" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1032 | apply (simp (no_asm) add: zmod_def divAlg_def) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1033 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1034 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1035 | lemma zdiv_minus1: "#0 $< b ==> #-1 zdiv b = #-1" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1036 | apply (simp (no_asm_simp) add: zdiv_def divAlg_def) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1037 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1038 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1039 | lemma zmod_minus1: "#0 $< b ==> #-1 zmod b = b $- #1" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1040 | apply (simp (no_asm_simp) add: zmod_def divAlg_def) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1041 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1042 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1043 | (** a positive, b positive **) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1044 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1045 | lemma zdiv_pos_pos: "[| #0 $< a; #0 $<= b |] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1046 | ==> a zdiv b = fst (posDivAlg(<intify(a), intify(b)>))" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1047 | apply (simp (no_asm_simp) add: zdiv_def divAlg_def) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1048 | apply (auto simp add: zle_def) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1049 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1050 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1051 | lemma zmod_pos_pos: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1052 | "[| #0 $< a; #0 $<= b |] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1053 | ==> a zmod b = snd (posDivAlg(<intify(a), intify(b)>))" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1054 | apply (simp (no_asm_simp) add: zmod_def divAlg_def) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1055 | apply (auto simp add: zle_def) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1056 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1057 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1058 | (** a negative, b positive **) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1059 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1060 | lemma zdiv_neg_pos: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1061 | "[| a $< #0; #0 $< b |] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1062 | ==> a zdiv b = fst (negDivAlg(<intify(a), intify(b)>))" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1063 | apply (simp (no_asm_simp) add: zdiv_def divAlg_def) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1064 | apply (blast dest: zle_zless_trans) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1065 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1066 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1067 | lemma zmod_neg_pos: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1068 | "[| a $< #0; #0 $< b |] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1069 | ==> a zmod b = snd (negDivAlg(<intify(a), intify(b)>))" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1070 | apply (simp (no_asm_simp) add: zmod_def divAlg_def) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1071 | apply (blast dest: zle_zless_trans) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1072 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1073 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1074 | (** a positive, b negative **) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1075 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1076 | lemma zdiv_pos_neg: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1077 | "[| #0 $< a; b $< #0 |] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1078 | ==> a zdiv b = fst (negateSnd(negDivAlg (<$-a, $-b>)))" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1079 | apply (simp (no_asm_simp) add: zdiv_def divAlg_def intify_eq_0_iff_zle) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1080 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1081 | apply (blast dest: zle_zless_trans)+ | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1082 | apply (blast dest: zless_trans) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1083 | apply (blast intro: zless_imp_zle) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1084 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1085 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1086 | lemma zmod_pos_neg: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1087 | "[| #0 $< a; b $< #0 |] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1088 | ==> a zmod b = snd (negateSnd(negDivAlg (<$-a, $-b>)))" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1089 | apply (simp (no_asm_simp) add: zmod_def divAlg_def intify_eq_0_iff_zle) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1090 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1091 | apply (blast dest: zle_zless_trans)+ | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1092 | apply (blast dest: zless_trans) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1093 | apply (blast intro: zless_imp_zle) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1094 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1095 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1096 | (** a negative, b negative **) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1097 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1098 | lemma zdiv_neg_neg: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1099 | "[| a $< #0; b $<= #0 |] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1100 | ==> a zdiv b = fst (negateSnd(posDivAlg(<$-a, $-b>)))" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1101 | apply (simp (no_asm_simp) add: zdiv_def divAlg_def) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1102 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1103 | apply (blast dest!: zle_zless_trans)+ | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1104 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1105 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1106 | lemma zmod_neg_neg: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1107 | "[| a $< #0; b $<= #0 |] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1108 | ==> a zmod b = snd (negateSnd(posDivAlg(<$-a, $-b>)))" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1109 | apply (simp (no_asm_simp) add: zmod_def divAlg_def) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1110 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1111 | apply (blast dest!: zle_zless_trans)+ | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1112 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1113 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1114 | declare zdiv_pos_pos [of "integ_of (v)" "integ_of (w)", standard, simp] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1115 | declare zdiv_neg_pos [of "integ_of (v)" "integ_of (w)", standard, simp] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1116 | declare zdiv_pos_neg [of "integ_of (v)" "integ_of (w)", standard, simp] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1117 | declare zdiv_neg_neg [of "integ_of (v)" "integ_of (w)", standard, simp] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1118 | declare zmod_pos_pos [of "integ_of (v)" "integ_of (w)", standard, simp] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1119 | declare zmod_neg_pos [of "integ_of (v)" "integ_of (w)", standard, simp] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1120 | declare zmod_pos_neg [of "integ_of (v)" "integ_of (w)", standard, simp] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1121 | declare zmod_neg_neg [of "integ_of (v)" "integ_of (w)", standard, simp] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1122 | declare posDivAlg_eqn [of concl: "integ_of (v)" "integ_of (w)", standard, simp] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1123 | declare negDivAlg_eqn [of concl: "integ_of (v)" "integ_of (w)", standard, simp] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1124 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1125 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1126 | (** Special-case simplification **) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1127 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1128 | lemma zmod_1 [simp]: "a zmod #1 = #0" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1129 | apply (cut_tac a = "a" and b = "#1" in pos_mod_sign) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1130 | apply (cut_tac [2] a = "a" and b = "#1" in pos_mod_bound) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1131 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1132 | (*arithmetic*) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1133 | apply (drule add1_zle_iff [THEN iffD2]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1134 | apply (rule zle_anti_sym) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1135 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1136 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1137 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1138 | lemma zdiv_1 [simp]: "a zdiv #1 = intify(a)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1139 | apply (cut_tac a = "a" and b = "#1" in zmod_zdiv_equality) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1140 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1141 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1142 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1143 | lemma zmod_minus1_right [simp]: "a zmod #-1 = #0" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1144 | apply (cut_tac a = "a" and b = "#-1" in neg_mod_sign) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1145 | apply (cut_tac [2] a = "a" and b = "#-1" in neg_mod_bound) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1146 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1147 | (*arithmetic*) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1148 | apply (drule add1_zle_iff [THEN iffD2]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1149 | apply (rule zle_anti_sym) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1150 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1151 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1152 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1153 | lemma zdiv_minus1_right_raw: "a \<in> int ==> a zdiv #-1 = $-a" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1154 | apply (cut_tac a = "a" and b = "#-1" in zmod_zdiv_equality) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1155 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1156 | apply (rule equation_zminus [THEN iffD2]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1157 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1158 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1159 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1160 | lemma zdiv_minus1_right: "a zdiv #-1 = $-a" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1161 | apply (cut_tac a = "intify (a)" in zdiv_minus1_right_raw) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1162 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1163 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1164 | declare zdiv_minus1_right [simp] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1165 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1166 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1167 | subsection{* Monotonicity in the first argument (divisor) *}
 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1168 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1169 | lemma zdiv_mono1: "[| a $<= a'; #0 $< b |] ==> a zdiv b $<= a' zdiv b" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1170 | apply (cut_tac a = "a" and b = "b" in zmod_zdiv_equality) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1171 | apply (cut_tac a = "a'" and b = "b" in zmod_zdiv_equality) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1172 | apply (rule unique_quotient_lemma) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1173 | apply (erule subst) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1174 | apply (erule subst) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1175 | apply (simp_all (no_asm_simp) add: pos_mod_sign pos_mod_bound) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1176 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1177 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1178 | lemma zdiv_mono1_neg: "[| a $<= a'; b $< #0 |] ==> a' zdiv b $<= a zdiv b" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1179 | apply (cut_tac a = "a" and b = "b" in zmod_zdiv_equality) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1180 | apply (cut_tac a = "a'" and b = "b" in zmod_zdiv_equality) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1181 | apply (rule unique_quotient_lemma_neg) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1182 | apply (erule subst) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1183 | apply (erule subst) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1184 | apply (simp_all (no_asm_simp) add: neg_mod_sign neg_mod_bound) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1185 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1186 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1187 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1188 | subsection{* Monotonicity in the second argument (dividend) *}
 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1189 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1190 | lemma q_pos_lemma: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1191 | "[| #0 $<= b'$*q' $+ r'; r' $< b'; #0 $< b' |] ==> #0 $<= q'" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1192 | apply (subgoal_tac "#0 $< b'$* (q' $+ #1)") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1193 | apply (simp add: int_0_less_mult_iff) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1194 | apply (blast dest: zless_trans intro: zless_add1_iff_zle [THEN iffD1]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1195 | apply (simp add: zadd_zmult_distrib2) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1196 | apply (erule zle_zless_trans) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1197 | apply (erule zadd_zless_mono2) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1198 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1199 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1200 | lemma zdiv_mono2_lemma: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1201 | "[| b$*q $+ r = b'$*q' $+ r'; #0 $<= b'$*q' $+ r'; | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1202 | r' $< b'; #0 $<= r; #0 $< b'; b' $<= b |] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1203 | ==> q $<= q'" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1204 | apply (frule q_pos_lemma, assumption+) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1205 | apply (subgoal_tac "b$*q $< b$* (q' $+ #1)") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1206 | apply (simp add: zmult_zless_cancel1) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1207 | apply (force dest: zless_add1_iff_zle [THEN iffD1] zless_trans zless_zle_trans) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1208 | apply (subgoal_tac "b$*q = r' $- r $+ b'$*q'") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1209 | prefer 2 apply (simp add: zcompare_rls) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1210 | apply (simp (no_asm_simp) add: zadd_zmult_distrib2) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1211 | apply (subst zadd_commute [of "b $\<times> q'"], rule zadd_zless_mono) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1212 | prefer 2 apply (blast intro: zmult_zle_mono1) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1213 | apply (subgoal_tac "r' $+ #0 $< b $+ r") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1214 | apply (simp add: zcompare_rls) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1215 | apply (rule zadd_zless_mono) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1216 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1217 | apply (blast dest: zless_zle_trans) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1218 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1219 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1220 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1221 | lemma zdiv_mono2_raw: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1222 | "[| #0 $<= a; #0 $< b'; b' $<= b; a \<in> int |] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1223 | ==> a zdiv b $<= a zdiv b'" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1224 | apply (subgoal_tac "#0 $< b") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1225 | prefer 2 apply (blast dest: zless_zle_trans) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1226 | apply (cut_tac a = "a" and b = "b" in zmod_zdiv_equality) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1227 | apply (cut_tac a = "a" and b = "b'" in zmod_zdiv_equality) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1228 | apply (rule zdiv_mono2_lemma) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1229 | apply (erule subst) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1230 | apply (erule subst) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1231 | apply (simp_all add: pos_mod_sign pos_mod_bound) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1232 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1233 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1234 | lemma zdiv_mono2: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1235 | "[| #0 $<= a; #0 $< b'; b' $<= b |] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1236 | ==> a zdiv b $<= a zdiv b'" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1237 | apply (cut_tac a = "intify (a)" in zdiv_mono2_raw) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1238 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1239 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1240 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1241 | lemma q_neg_lemma: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1242 | "[| b'$*q' $+ r' $< #0; #0 $<= r'; #0 $< b' |] ==> q' $< #0" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1243 | apply (subgoal_tac "b'$*q' $< #0") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1244 | prefer 2 apply (force intro: zle_zless_trans) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1245 | apply (simp add: zmult_less_0_iff) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1246 | apply (blast dest: zless_trans) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1247 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1248 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1249 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1250 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1251 | lemma zdiv_mono2_neg_lemma: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1252 | "[| b$*q $+ r = b'$*q' $+ r'; b'$*q' $+ r' $< #0; | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1253 | r $< b; #0 $<= r'; #0 $< b'; b' $<= b |] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1254 | ==> q' $<= q" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1255 | apply (subgoal_tac "#0 $< b") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1256 | prefer 2 apply (blast dest: zless_zle_trans) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1257 | apply (frule q_neg_lemma, assumption+) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1258 | apply (subgoal_tac "b$*q' $< b$* (q $+ #1)") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1259 | apply (simp add: zmult_zless_cancel1) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1260 | apply (blast dest: zless_trans zless_add1_iff_zle [THEN iffD1]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1261 | apply (simp (no_asm_simp) add: zadd_zmult_distrib2) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1262 | apply (subgoal_tac "b$*q' $<= b'$*q'") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1263 | prefer 2 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1264 | apply (simp add: zmult_zle_cancel2) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1265 | apply (blast dest: zless_trans) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1266 | apply (subgoal_tac "b'$*q' $+ r $< b $+ (b$*q $+ r)") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1267 | prefer 2 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1268 | apply (erule ssubst) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1269 | apply simp | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1270 | apply (drule_tac w' = "r" and z' = "#0" in zadd_zless_mono) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1271 | apply (assumption) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1272 | apply simp | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1273 | apply (simp (no_asm_use) add: zadd_commute) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1274 | apply (rule zle_zless_trans) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1275 | prefer 2 apply (assumption) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1276 | apply (simp (no_asm_simp) add: zmult_zle_cancel2) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1277 | apply (blast dest: zless_trans) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1278 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1279 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1280 | lemma zdiv_mono2_neg_raw: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1281 | "[| a $< #0; #0 $< b'; b' $<= b; a \<in> int |] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1282 | ==> a zdiv b' $<= a zdiv b" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1283 | apply (subgoal_tac "#0 $< b") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1284 | prefer 2 apply (blast dest: zless_zle_trans) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1285 | apply (cut_tac a = "a" and b = "b" in zmod_zdiv_equality) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1286 | apply (cut_tac a = "a" and b = "b'" in zmod_zdiv_equality) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1287 | apply (rule zdiv_mono2_neg_lemma) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1288 | apply (erule subst) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1289 | apply (erule subst) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1290 | apply (simp_all add: pos_mod_sign pos_mod_bound) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1291 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1292 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1293 | lemma zdiv_mono2_neg: "[| a $< #0; #0 $< b'; b' $<= b |] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1294 | ==> a zdiv b' $<= a zdiv b" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1295 | apply (cut_tac a = "intify (a)" in zdiv_mono2_neg_raw) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1296 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1297 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1298 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1299 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1300 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1301 | subsection{* More algebraic laws for zdiv and zmod *}
 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1302 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1303 | (** proving (a*b) zdiv c = a $* (b zdiv c) $+ a * (b zmod c) **) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1304 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1305 | lemma zmult1_lemma: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1306 | "[| quorem(<b,c>, <q,r>); c \<in> int; c \<noteq> #0 |] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1307 | ==> quorem (<a$*b, c>, <a$*q $+ (a$*r) zdiv c, (a$*r) zmod c>)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1308 | apply (auto simp add: split_ifs quorem_def neq_iff_zless zadd_zmult_distrib2 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1309 | pos_mod_sign pos_mod_bound neg_mod_sign neg_mod_bound) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1310 | apply (auto intro: raw_zmod_zdiv_equality) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1311 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1312 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1313 | lemma zdiv_zmult1_eq_raw: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1314 | "[|b \<in> int; c \<in> int|] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1315 | ==> (a$*b) zdiv c = a$*(b zdiv c) $+ a$*(b zmod c) zdiv c" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1316 | apply (case_tac "c = #0") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1317 | apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1318 | apply (rule quorem_div_mod [THEN zmult1_lemma, THEN quorem_div]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1319 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1320 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1321 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1322 | lemma zdiv_zmult1_eq: "(a$*b) zdiv c = a$*(b zdiv c) $+ a$*(b zmod c) zdiv c" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1323 | apply (cut_tac b = "intify (b)" and c = "intify (c)" in zdiv_zmult1_eq_raw) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1324 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1325 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1326 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1327 | lemma zmod_zmult1_eq_raw: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1328 | "[|b \<in> int; c \<in> int|] ==> (a$*b) zmod c = a$*(b zmod c) zmod c" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1329 | apply (case_tac "c = #0") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1330 | apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1331 | apply (rule quorem_div_mod [THEN zmult1_lemma, THEN quorem_mod]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1332 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1333 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1334 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1335 | lemma zmod_zmult1_eq: "(a$*b) zmod c = a$*(b zmod c) zmod c" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1336 | apply (cut_tac b = "intify (b)" and c = "intify (c)" in zmod_zmult1_eq_raw) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1337 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1338 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1339 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1340 | lemma zmod_zmult1_eq': "(a$*b) zmod c = ((a zmod c) $* b) zmod c" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1341 | apply (rule trans) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1342 | apply (rule_tac b = " (b $* a) zmod c" in trans) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1343 | apply (rule_tac [2] zmod_zmult1_eq) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1344 | apply (simp_all (no_asm) add: zmult_commute) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1345 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1346 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1347 | lemma zmod_zmult_distrib: "(a$*b) zmod c = ((a zmod c) $* (b zmod c)) zmod c" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1348 | apply (rule zmod_zmult1_eq' [THEN trans]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1349 | apply (rule zmod_zmult1_eq) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1350 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1351 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1352 | lemma zdiv_zmult_self1 [simp]: "intify(b) \<noteq> #0 ==> (a$*b) zdiv b = intify(a)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1353 | apply (simp (no_asm_simp) add: zdiv_zmult1_eq) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1354 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1355 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1356 | lemma zdiv_zmult_self2 [simp]: "intify(b) \<noteq> #0 ==> (b$*a) zdiv b = intify(a)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1357 | apply (subst zmult_commute , erule zdiv_zmult_self1) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1358 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1359 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1360 | lemma zmod_zmult_self1 [simp]: "(a$*b) zmod b = #0" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1361 | apply (simp (no_asm) add: zmod_zmult1_eq) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1362 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1363 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1364 | lemma zmod_zmult_self2 [simp]: "(b$*a) zmod b = #0" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1365 | apply (simp (no_asm) add: zmult_commute zmod_zmult1_eq) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1366 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1367 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1368 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1369 | (** proving (a$+b) zdiv c = | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1370 | a zdiv c $+ b zdiv c $+ ((a zmod c $+ b zmod c) zdiv c) **) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1371 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1372 | lemma zadd1_lemma: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1373 | "[| quorem(<a,c>, <aq,ar>); quorem(<b,c>, <bq,br>); | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1374 | c \<in> int; c \<noteq> #0 |] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1375 | ==> quorem (<a$+b, c>, <aq $+ bq $+ (ar$+br) zdiv c, (ar$+br) zmod c>)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1376 | apply (auto simp add: split_ifs quorem_def neq_iff_zless zadd_zmult_distrib2 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1377 | pos_mod_sign pos_mod_bound neg_mod_sign neg_mod_bound) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1378 | apply (auto intro: raw_zmod_zdiv_equality) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1379 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1380 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1381 | (*NOT suitable for rewriting: the RHS has an instance of the LHS*) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1382 | lemma zdiv_zadd1_eq_raw: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1383 | "[|a \<in> int; b \<in> int; c \<in> int|] ==> | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1384 | (a$+b) zdiv c = a zdiv c $+ b zdiv c $+ ((a zmod c $+ b zmod c) zdiv c)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1385 | apply (case_tac "c = #0") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1386 | apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1387 | apply (blast intro: zadd1_lemma [OF quorem_div_mod quorem_div_mod, | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1388 | THEN quorem_div]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1389 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1390 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1391 | lemma zdiv_zadd1_eq: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1392 | "(a$+b) zdiv c = a zdiv c $+ b zdiv c $+ ((a zmod c $+ b zmod c) zdiv c)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1393 | apply (cut_tac a = "intify (a)" and b = "intify (b)" and c = "intify (c)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1394 | in zdiv_zadd1_eq_raw) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1395 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1396 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1397 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1398 | lemma zmod_zadd1_eq_raw: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1399 | "[|a \<in> int; b \<in> int; c \<in> int|] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1400 | ==> (a$+b) zmod c = (a zmod c $+ b zmod c) zmod c" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1401 | apply (case_tac "c = #0") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1402 | apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1403 | apply (blast intro: zadd1_lemma [OF quorem_div_mod quorem_div_mod, | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1404 | THEN quorem_mod]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1405 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1406 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1407 | lemma zmod_zadd1_eq: "(a$+b) zmod c = (a zmod c $+ b zmod c) zmod c" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1408 | apply (cut_tac a = "intify (a)" and b = "intify (b)" and c = "intify (c)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1409 | in zmod_zadd1_eq_raw) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1410 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1411 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1412 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1413 | lemma zmod_div_trivial_raw: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1414 | "[|a \<in> int; b \<in> int|] ==> (a zmod b) zdiv b = #0" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1415 | apply (case_tac "b = #0") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1416 | apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1417 | apply (auto simp add: neq_iff_zless pos_mod_sign pos_mod_bound | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1418 | zdiv_pos_pos_trivial neg_mod_sign neg_mod_bound zdiv_neg_neg_trivial) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1419 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1420 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1421 | lemma zmod_div_trivial [simp]: "(a zmod b) zdiv b = #0" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1422 | apply (cut_tac a = "intify (a)" and b = "intify (b)" in zmod_div_trivial_raw) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1423 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1424 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1425 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1426 | lemma zmod_mod_trivial_raw: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1427 | "[|a \<in> int; b \<in> int|] ==> (a zmod b) zmod b = a zmod b" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1428 | apply (case_tac "b = #0") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1429 | apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1430 | apply (auto simp add: neq_iff_zless pos_mod_sign pos_mod_bound | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1431 | zmod_pos_pos_trivial neg_mod_sign neg_mod_bound zmod_neg_neg_trivial) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1432 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1433 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1434 | lemma zmod_mod_trivial [simp]: "(a zmod b) zmod b = a zmod b" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1435 | apply (cut_tac a = "intify (a)" and b = "intify (b)" in zmod_mod_trivial_raw) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1436 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1437 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1438 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1439 | lemma zmod_zadd_left_eq: "(a$+b) zmod c = ((a zmod c) $+ b) zmod c" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1440 | apply (rule trans [symmetric]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1441 | apply (rule zmod_zadd1_eq) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1442 | apply (simp (no_asm)) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1443 | apply (rule zmod_zadd1_eq [symmetric]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1444 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1445 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1446 | lemma zmod_zadd_right_eq: "(a$+b) zmod c = (a $+ (b zmod c)) zmod c" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1447 | apply (rule trans [symmetric]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1448 | apply (rule zmod_zadd1_eq) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1449 | apply (simp (no_asm)) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1450 | apply (rule zmod_zadd1_eq [symmetric]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1451 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1452 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1453 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1454 | lemma zdiv_zadd_self1 [simp]: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1455 | "intify(a) \<noteq> #0 ==> (a$+b) zdiv a = b zdiv a $+ #1" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1456 | by (simp (no_asm_simp) add: zdiv_zadd1_eq) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1457 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1458 | lemma zdiv_zadd_self2 [simp]: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1459 | "intify(a) \<noteq> #0 ==> (b$+a) zdiv a = b zdiv a $+ #1" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1460 | by (simp (no_asm_simp) add: zdiv_zadd1_eq) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1461 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1462 | lemma zmod_zadd_self1 [simp]: "(a$+b) zmod a = b zmod a" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1463 | apply (case_tac "a = #0") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1464 | apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1465 | apply (simp (no_asm_simp) add: zmod_zadd1_eq) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1466 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1467 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1468 | lemma zmod_zadd_self2 [simp]: "(b$+a) zmod a = b zmod a" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1469 | apply (case_tac "a = #0") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1470 | apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1471 | apply (simp (no_asm_simp) add: zmod_zadd1_eq) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1472 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1473 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1474 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1475 | subsection{* proving  a zdiv (b*c) = (a zdiv b) zdiv c *}
 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1476 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1477 | (*The condition c>0 seems necessary. Consider that 7 zdiv ~6 = ~2 but | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1478 | 7 zdiv 2 zdiv ~3 = 3 zdiv ~3 = ~1. The subcase (a zdiv b) zmod c = 0 seems | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1479 | to cause particular problems.*) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1480 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1481 | (** first, four lemmas to bound the remainder for the cases b<0 and b>0 **) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1482 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1483 | lemma zdiv_zmult2_aux1: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1484 | "[| #0 $< c; b $< r; r $<= #0 |] ==> b$*c $< b$*(q zmod c) $+ r" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1485 | apply (subgoal_tac "b $* (c $- q zmod c) $< r $* #1") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1486 | apply (simp add: zdiff_zmult_distrib2 zadd_commute zcompare_rls) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1487 | apply (rule zle_zless_trans) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1488 | apply (erule_tac [2] zmult_zless_mono1) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1489 | apply (rule zmult_zle_mono2_neg) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1490 | apply (auto simp add: zcompare_rls zadd_commute add1_zle_iff pos_mod_bound) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1491 | apply (blast intro: zless_imp_zle dest: zless_zle_trans) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1492 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1493 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1494 | lemma zdiv_zmult2_aux2: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1495 | "[| #0 $< c; b $< r; r $<= #0 |] ==> b $* (q zmod c) $+ r $<= #0" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1496 | apply (subgoal_tac "b $* (q zmod c) $<= #0") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1497 | prefer 2 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1498 | apply (simp add: zmult_le_0_iff pos_mod_sign) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1499 | apply (blast intro: zless_imp_zle dest: zless_zle_trans) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1500 | (*arithmetic*) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1501 | apply (drule zadd_zle_mono) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1502 | apply assumption | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1503 | apply (simp add: zadd_commute) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1504 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1505 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1506 | lemma zdiv_zmult2_aux3: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1507 | "[| #0 $< c; #0 $<= r; r $< b |] ==> #0 $<= b $* (q zmod c) $+ r" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1508 | apply (subgoal_tac "#0 $<= b $* (q zmod c)") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1509 | prefer 2 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1510 | apply (simp add: int_0_le_mult_iff pos_mod_sign) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1511 | apply (blast intro: zless_imp_zle dest: zle_zless_trans) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1512 | (*arithmetic*) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1513 | apply (drule zadd_zle_mono) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1514 | apply assumption | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1515 | apply (simp add: zadd_commute) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1516 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1517 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1518 | lemma zdiv_zmult2_aux4: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1519 | "[| #0 $< c; #0 $<= r; r $< b |] ==> b $* (q zmod c) $+ r $< b $* c" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1520 | apply (subgoal_tac "r $* #1 $< b $* (c $- q zmod c)") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1521 | apply (simp add: zdiff_zmult_distrib2 zadd_commute zcompare_rls) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1522 | apply (rule zless_zle_trans) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1523 | apply (erule zmult_zless_mono1) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1524 | apply (rule_tac [2] zmult_zle_mono2) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1525 | apply (auto simp add: zcompare_rls zadd_commute add1_zle_iff pos_mod_bound) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1526 | apply (blast intro: zless_imp_zle dest: zle_zless_trans) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1527 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1528 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1529 | lemma zdiv_zmult2_lemma: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1530 | "[| quorem (<a,b>, <q,r>); a \<in> int; b \<in> int; b \<noteq> #0; #0 $< c |] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1531 | ==> quorem (<a,b$*c>, <q zdiv c, b$*(q zmod c) $+ r>)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1532 | apply (auto simp add: zmult_ac zmod_zdiv_equality [symmetric] quorem_def | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1533 | neq_iff_zless int_0_less_mult_iff | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1534 | zadd_zmult_distrib2 [symmetric] zdiv_zmult2_aux1 zdiv_zmult2_aux2 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1535 | zdiv_zmult2_aux3 zdiv_zmult2_aux4) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1536 | apply (blast dest: zless_trans)+ | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1537 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1538 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1539 | lemma zdiv_zmult2_eq_raw: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1540 | "[|#0 $< c; a \<in> int; b \<in> int|] ==> a zdiv (b$*c) = (a zdiv b) zdiv c" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1541 | apply (case_tac "b = #0") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1542 | apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1543 | apply (rule quorem_div_mod [THEN zdiv_zmult2_lemma, THEN quorem_div]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1544 | apply (auto simp add: intify_eq_0_iff_zle) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1545 | apply (blast dest: zle_zless_trans) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1546 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1547 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1548 | lemma zdiv_zmult2_eq: "#0 $< c ==> a zdiv (b$*c) = (a zdiv b) zdiv c" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1549 | apply (cut_tac a = "intify (a)" and b = "intify (b)" in zdiv_zmult2_eq_raw) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1550 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1551 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1552 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1553 | lemma zmod_zmult2_eq_raw: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1554 | "[|#0 $< c; a \<in> int; b \<in> int|] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1555 | ==> a zmod (b$*c) = b$*(a zdiv b zmod c) $+ a zmod b" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1556 | apply (case_tac "b = #0") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1557 | apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1558 | apply (rule quorem_div_mod [THEN zdiv_zmult2_lemma, THEN quorem_mod]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1559 | apply (auto simp add: intify_eq_0_iff_zle) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1560 | apply (blast dest: zle_zless_trans) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1561 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1562 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1563 | lemma zmod_zmult2_eq: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1564 | "#0 $< c ==> a zmod (b$*c) = b$*(a zdiv b zmod c) $+ a zmod b" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1565 | apply (cut_tac a = "intify (a)" and b = "intify (b)" in zmod_zmult2_eq_raw) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1566 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1567 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1568 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1569 | subsection{* Cancellation of common factors in "zdiv" *}
 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1570 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1571 | lemma zdiv_zmult_zmult1_aux1: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1572 | "[| #0 $< b; intify(c) \<noteq> #0 |] ==> (c$*a) zdiv (c$*b) = a zdiv b" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1573 | apply (subst zdiv_zmult2_eq) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1574 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1575 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1576 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1577 | lemma zdiv_zmult_zmult1_aux2: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1578 | "[| b $< #0; intify(c) \<noteq> #0 |] ==> (c$*a) zdiv (c$*b) = a zdiv b" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1579 | apply (subgoal_tac " (c $* ($-a)) zdiv (c $* ($-b)) = ($-a) zdiv ($-b)") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1580 | apply (rule_tac [2] zdiv_zmult_zmult1_aux1) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1581 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1582 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1583 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1584 | lemma zdiv_zmult_zmult1_raw: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1585 | "[|intify(c) \<noteq> #0; b \<in> int|] ==> (c$*a) zdiv (c$*b) = a zdiv b" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1586 | apply (case_tac "b = #0") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1587 | apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1588 | apply (auto simp add: neq_iff_zless [of b] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1589 | zdiv_zmult_zmult1_aux1 zdiv_zmult_zmult1_aux2) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1590 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1591 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1592 | lemma zdiv_zmult_zmult1: "intify(c) \<noteq> #0 ==> (c$*a) zdiv (c$*b) = a zdiv b" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1593 | apply (cut_tac b = "intify (b)" in zdiv_zmult_zmult1_raw) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1594 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1595 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1596 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1597 | lemma zdiv_zmult_zmult2: "intify(c) \<noteq> #0 ==> (a$*c) zdiv (b$*c) = a zdiv b" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1598 | apply (drule zdiv_zmult_zmult1) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1599 | apply (auto simp add: zmult_commute) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1600 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1601 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1602 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1603 | subsection{* Distribution of factors over "zmod" *}
 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1604 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1605 | lemma zmod_zmult_zmult1_aux1: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1606 | "[| #0 $< b; intify(c) \<noteq> #0 |] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1607 | ==> (c$*a) zmod (c$*b) = c $* (a zmod b)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1608 | apply (subst zmod_zmult2_eq) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1609 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1610 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1611 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1612 | lemma zmod_zmult_zmult1_aux2: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1613 | "[| b $< #0; intify(c) \<noteq> #0 |] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1614 | ==> (c$*a) zmod (c$*b) = c $* (a zmod b)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1615 | apply (subgoal_tac " (c $* ($-a)) zmod (c $* ($-b)) = c $* (($-a) zmod ($-b))") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1616 | apply (rule_tac [2] zmod_zmult_zmult1_aux1) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1617 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1618 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1619 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1620 | lemma zmod_zmult_zmult1_raw: | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1621 | "[|b \<in> int; c \<in> int|] ==> (c$*a) zmod (c$*b) = c $* (a zmod b)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1622 | apply (case_tac "b = #0") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1623 | apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1624 | apply (case_tac "c = #0") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1625 | apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1626 | apply (auto simp add: neq_iff_zless [of b] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1627 | zmod_zmult_zmult1_aux1 zmod_zmult_zmult1_aux2) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1628 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1629 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1630 | lemma zmod_zmult_zmult1: "(c$*a) zmod (c$*b) = c $* (a zmod b)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1631 | apply (cut_tac b = "intify (b)" and c = "intify (c)" in zmod_zmult_zmult1_raw) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1632 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1633 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1634 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1635 | lemma zmod_zmult_zmult2: "(a$*c) zmod (b$*c) = (a zmod b) $* c" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1636 | apply (cut_tac c = "c" in zmod_zmult_zmult1) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1637 | apply (auto simp add: zmult_commute) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1638 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1639 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1640 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1641 | (** Quotients of signs **) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1642 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1643 | lemma zdiv_neg_pos_less0: "[| a $< #0; #0 $< b |] ==> a zdiv b $< #0" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1644 | apply (subgoal_tac "a zdiv b $<= #-1") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1645 | apply (erule zle_zless_trans) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1646 | apply (simp (no_asm)) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1647 | apply (rule zle_trans) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1648 | apply (rule_tac a' = "#-1" in zdiv_mono1) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1649 | apply (rule zless_add1_iff_zle [THEN iffD1]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1650 | apply (simp (no_asm)) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1651 | apply (auto simp add: zdiv_minus1) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1652 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1653 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1654 | lemma zdiv_nonneg_neg_le0: "[| #0 $<= a; b $< #0 |] ==> a zdiv b $<= #0" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1655 | apply (drule zdiv_mono1_neg) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1656 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1657 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1658 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1659 | lemma pos_imp_zdiv_nonneg_iff: "#0 $< b ==> (#0 $<= a zdiv b) <-> (#0 $<= a)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1660 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1661 | apply (drule_tac [2] zdiv_mono1) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1662 | apply (auto simp add: neq_iff_zless) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1663 | apply (simp (no_asm_use) add: not_zless_iff_zle [THEN iff_sym]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1664 | apply (blast intro: zdiv_neg_pos_less0) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1665 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1666 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1667 | lemma neg_imp_zdiv_nonneg_iff: "b $< #0 ==> (#0 $<= a zdiv b) <-> (a $<= #0)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1668 | apply (subst zdiv_zminus_zminus [symmetric]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1669 | apply (rule iff_trans) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1670 | apply (rule pos_imp_zdiv_nonneg_iff) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1671 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1672 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1673 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1674 | (*But not (a zdiv b $<= 0 iff a$<=0); consider a=1, b=2 when a zdiv b = 0.*) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1675 | lemma pos_imp_zdiv_neg_iff: "#0 $< b ==> (a zdiv b $< #0) <-> (a $< #0)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1676 | apply (simp (no_asm_simp) add: not_zle_iff_zless [THEN iff_sym]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1677 | apply (erule pos_imp_zdiv_nonneg_iff) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1678 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1679 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1680 | (*Again the law fails for $<=: consider a = -1, b = -2 when a zdiv b = 0*) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1681 | lemma neg_imp_zdiv_neg_iff: "b $< #0 ==> (a zdiv b $< #0) <-> (#0 $< a)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1682 | apply (simp (no_asm_simp) add: not_zle_iff_zless [THEN iff_sym]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1683 | apply (erule neg_imp_zdiv_nonneg_iff) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1684 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1685 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1686 | (* | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1687 | THESE REMAIN TO BE CONVERTED -- but aren't that useful! | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1688 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1689 |  subsection{* Speeding up the division algorithm with shifting *}
 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1690 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1691 | (** computing "zdiv" by shifting **) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1692 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1693 | lemma pos_zdiv_mult_2: "#0 $<= a ==> (#1 $+ #2$*b) zdiv (#2$*a) = b zdiv a" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1694 | apply (case_tac "a = #0") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1695 | apply (subgoal_tac "#1 $<= a") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1696 | apply (arith_tac 2) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1697 | apply (subgoal_tac "#1 $< a $* #2") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1698 | apply (arith_tac 2) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1699 | apply (subgoal_tac "#2$* (#1 $+ b zmod a) $<= #2$*a") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1700 | apply (rule_tac [2] zmult_zle_mono2) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1701 | apply (auto simp add: zadd_commute zmult_commute add1_zle_iff pos_mod_bound) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1702 | apply (subst zdiv_zadd1_eq) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1703 | apply (simp (no_asm_simp) add: zdiv_zmult_zmult2 zmod_zmult_zmult2 zdiv_pos_pos_trivial) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1704 | apply (subst zdiv_pos_pos_trivial) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1705 | apply (simp (no_asm_simp) add: [zmod_pos_pos_trivial pos_mod_sign [THEN zadd_zle_mono1] RSN (2,zle_trans) ]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1706 | apply (auto simp add: zmod_pos_pos_trivial) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1707 | apply (subgoal_tac "#0 $<= b zmod a") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1708 | apply (asm_simp_tac (simpset () add: pos_mod_sign) 2) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1709 | apply arith | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1710 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1711 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1712 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1713 | lemma neg_zdiv_mult_2: "a $<= #0 ==> (#1 $+ #2$*b) zdiv (#2$*a) <-> (b$+#1) zdiv a" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1714 | apply (subgoal_tac " (#1 $+ #2$* ($-b-#1)) zdiv (#2 $* ($-a)) <-> ($-b-#1) zdiv ($-a)") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1715 | apply (rule_tac [2] pos_zdiv_mult_2) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1716 | apply (auto simp add: zmult_zminus_right) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1717 | apply (subgoal_tac " (#-1 - (#2 $* b)) = - (#1 $+ (#2 $* b))") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1718 | apply (Simp_tac 2) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1719 | apply (asm_full_simp_tac (HOL_ss add: zdiv_zminus_zminus zdiff_def zminus_zadd_distrib [symmetric]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1720 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1721 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1722 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1723 | (*Not clear why this must be proved separately; probably integ_of causes | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1724 | simplification problems*) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1725 | lemma lemma: "~ #0 $<= x ==> x $<= #0" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1726 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1727 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1728 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1729 | lemma zdiv_integ_of_BIT: "integ_of (v BIT b) zdiv integ_of (w BIT False) = | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1730 | (if ~b | #0 $<= integ_of w | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1731 | then integ_of v zdiv (integ_of w) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1732 | else (integ_of v $+ #1) zdiv (integ_of w))" | 
| 32149 
ef59550a55d3
renamed simpset_of to global_simpset_of, and local_simpset_of to simpset_of -- same for claset and clasimpset;
 wenzelm parents: 
26056diff
changeset | 1733 | apply (simp_tac (global_simpset_of Int.thy add: zadd_assoc integ_of_BIT) | 
| 26056 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1734 | apply (simp (no_asm_simp) del: bin_arith_extra_simps@bin_rel_simps add: zdiv_zmult_zmult1 pos_zdiv_mult_2 lemma neg_zdiv_mult_2) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1735 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1736 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1737 | declare zdiv_integ_of_BIT [simp] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1738 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1739 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1740 | (** computing "zmod" by shifting (proofs resemble those for "zdiv") **) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1741 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1742 | lemma pos_zmod_mult_2: "#0 $<= a ==> (#1 $+ #2$*b) zmod (#2$*a) = #1 $+ #2 $* (b zmod a)" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1743 | apply (case_tac "a = #0") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1744 | apply (subgoal_tac "#1 $<= a") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1745 | apply (arith_tac 2) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1746 | apply (subgoal_tac "#1 $< a $* #2") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1747 | apply (arith_tac 2) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1748 | apply (subgoal_tac "#2$* (#1 $+ b zmod a) $<= #2$*a") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1749 | apply (rule_tac [2] zmult_zle_mono2) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1750 | apply (auto simp add: zadd_commute zmult_commute add1_zle_iff pos_mod_bound) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1751 | apply (subst zmod_zadd1_eq) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1752 | apply (simp (no_asm_simp) add: zmod_zmult_zmult2 zmod_pos_pos_trivial) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1753 | apply (rule zmod_pos_pos_trivial) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1754 | apply (simp (no_asm_simp) # add: [zmod_pos_pos_trivial pos_mod_sign [THEN zadd_zle_mono1] RSN (2,zle_trans) ]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1755 | apply (auto simp add: zmod_pos_pos_trivial) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1756 | apply (subgoal_tac "#0 $<= b zmod a") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1757 | apply (asm_simp_tac (simpset () add: pos_mod_sign) 2) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1758 | apply arith | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1759 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1760 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1761 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1762 | lemma neg_zmod_mult_2: "a $<= #0 ==> (#1 $+ #2$*b) zmod (#2$*a) = #2 $* ((b$+#1) zmod a) - #1" | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1763 | apply (subgoal_tac " (#1 $+ #2$* ($-b-#1)) zmod (#2$* ($-a)) = #1 $+ #2$* (($-b-#1) zmod ($-a))") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1764 | apply (rule_tac [2] pos_zmod_mult_2) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1765 | apply (auto simp add: zmult_zminus_right) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1766 | apply (subgoal_tac " (#-1 - (#2 $* b)) = - (#1 $+ (#2 $* b))") | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1767 | apply (Simp_tac 2) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1768 | apply (asm_full_simp_tac (HOL_ss add: zmod_zminus_zminus zdiff_def zminus_zadd_distrib [symmetric]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1769 | apply (dtac (zminus_equation [THEN iffD1, symmetric]) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1770 | apply auto | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1771 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1772 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1773 | lemma zmod_integ_of_BIT: "integ_of (v BIT b) zmod integ_of (w BIT False) = | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1774 | (if b then | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1775 | if #0 $<= integ_of w | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1776 | then #2 $* (integ_of v zmod integ_of w) $+ #1 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1777 | else #2 $* ((integ_of v $+ #1) zmod integ_of w) - #1 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1778 | else #2 $* (integ_of v zmod integ_of w))" | 
| 32149 
ef59550a55d3
renamed simpset_of to global_simpset_of, and local_simpset_of to simpset_of -- same for claset and clasimpset;
 wenzelm parents: 
26056diff
changeset | 1779 | apply (simp_tac (global_simpset_of Int.thy add: zadd_assoc integ_of_BIT) | 
| 26056 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1780 | apply (simp (no_asm_simp) del: bin_arith_extra_simps@bin_rel_simps add: zmod_zmult_zmult1 pos_zmod_mult_2 lemma neg_zmod_mult_2) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1781 | done | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1782 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1783 | declare zmod_integ_of_BIT [simp] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1784 | *) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1785 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1786 | end | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 1787 |