| author | wenzelm | 
| Thu, 03 Jan 2019 21:15:52 +0100 | |
| changeset 69587 | 53982d5ec0bb | 
| parent 68490 | eb53f944c8cd | 
| child 69593 | 3dda49e08b9d | 
| permissions | -rw-r--r-- | 
| 
68490
 
eb53f944c8cd
simplified ZF theory names (in contrast to 6a0801279f4c): session-qualification already achieves disjointness;
 
wenzelm 
parents: 
63648 
diff
changeset
 | 
1  | 
(* Title: ZF/IntDiv.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: 
32149 
diff
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: 
32149 
diff
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: 
32149 
diff
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: 
32149 
diff
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  | 
|
| 46820 | 21  | 
fun divAlg (a,b) = if 0<=a then  | 
22  | 
if b>0 then posDivAlg (a,b)  | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32149 
diff
changeset
 | 
23  | 
else if a=0 then (0,0)  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32149 
diff
changeset
 | 
24  | 
else negateSnd (negDivAlg (~a,~b))  | 
| 46820 | 25  | 
else  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32149 
diff
changeset
 | 
26  | 
if 0<b then negDivAlg (a,b)  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32149 
diff
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  | 
|
| 60770 | 30  | 
section\<open>The Division Operators Div and Mod\<close>  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
31  | 
|
| 
68490
 
eb53f944c8cd
simplified ZF theory names (in contrast to 6a0801279f4c): session-qualification already achieves disjointness;
 
wenzelm 
parents: 
63648 
diff
changeset
 | 
32  | 
theory IntDiv  | 
| 58022 | 33  | 
imports Bin OrderArith  | 
34  | 
begin  | 
|
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
35  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
36  | 
definition  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
37  | 
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
 | 
38  | 
"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
 | 
39  | 
a = b$*q $+ r &  | 
| 61395 | 40  | 
(#0$<b & #0$\<le>r & r$<b | ~(#0$<b) & b$<r & r $\<le> #0)"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
41  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
42  | 
definition  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
43  | 
adjust :: "[i,i] => i" where  | 
| 61395 | 44  | 
"adjust(b) == %<q,r>. if #0 $\<le> r$-b then <#2$*q $+ #1,r$-b>  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
45  | 
else <#2$*q,r>"  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
46  | 
|
| 
 
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  | 
(** the division algorithm **)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
49  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
50  | 
definition  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
51  | 
posDivAlg :: "i => i" where  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
52  | 
(*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
 | 
53  | 
(*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
 | 
54  | 
"posDivAlg(ab) ==  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
55  | 
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: 
32149 
diff
changeset
 | 
56  | 
ab,  | 
| 61395 | 57  | 
%<a,b> f. if (a$<b | b$\<le>#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
 | 
58  | 
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
 | 
59  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
60  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
61  | 
(*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
 | 
62  | 
definition  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
63  | 
negDivAlg :: "i => i" where  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
64  | 
(*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
 | 
65  | 
"negDivAlg(ab) ==  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
66  | 
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: 
32149 
diff
changeset
 | 
67  | 
ab,  | 
| 61395 | 68  | 
%<a,b> f. if (#0 $\<le> a$+b | b$\<le>#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
 | 
69  | 
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
 | 
70  | 
|
| 46820 | 71  | 
(*for the general case @{term"b\<noteq>0"}*)
 | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
72  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
73  | 
definition  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
74  | 
negateSnd :: "i => i" where  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
75  | 
"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
 | 
76  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
77  | 
(*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
 | 
78  | 
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
 | 
79  | 
definition  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
80  | 
divAlg :: "i => i" where  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
81  | 
"divAlg ==  | 
| 61395 | 82  | 
%<a,b>. if #0 $\<le> a then  | 
83  | 
if #0 $\<le> b then posDivAlg (<a,b>)  | 
|
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
84  | 
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
 | 
85  | 
else negateSnd (negDivAlg (<$-a,$-b>))  | 
| 46820 | 86  | 
else  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
87  | 
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
 | 
88  | 
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
 | 
89  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
90  | 
definition  | 
| 69587 | 91  | 
zdiv :: "[i,i]=>i" (infixl \<open>zdiv\<close> 70) where  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
92  | 
"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
 | 
93  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
94  | 
definition  | 
| 69587 | 95  | 
zmod :: "[i,i]=>i" (infixl \<open>zmod\<close> 70) where  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
96  | 
"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
 | 
97  | 
|
| 
 
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  | 
(** 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
 | 
100  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
101  | 
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
 | 
102  | 
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
 | 
103  | 
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
 | 
104  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
105  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
106  | 
|
| 61395 | 107  | 
lemma zpos_add_zpos_imp_zpos: "[| #0 $\<le> x; #0 $\<le> y |] ==> #0 $\<le> x $+ y"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
108  | 
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
 | 
109  | 
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
 | 
110  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
111  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
112  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
113  | 
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
 | 
114  | 
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
 | 
115  | 
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
 | 
116  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
117  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
118  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
119  | 
(* 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
 | 
120  | 
lemma zneg_or_0_add_zneg_or_0_imp_zneg_or_0:  | 
| 61395 | 121  | 
"[| x $\<le> #0; y $\<le> #0 |] ==> x $+ y $\<le> #0"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
122  | 
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
 | 
123  | 
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
 | 
124  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
125  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
126  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
127  | 
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
 | 
128  | 
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
 | 
129  | 
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
 | 
130  | 
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
 | 
131  | 
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
 | 
132  | 
apply simp  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
133  | 
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
 | 
134  | 
apply simp  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
135  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
136  | 
|
| 
 
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  | 
(*** 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
 | 
139  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
140  | 
lemma zless_add_succ_iff:  | 
| 
46821
 
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
 
paulson 
parents: 
46820 
diff
changeset
 | 
141  | 
"(w $< z $+ $# succ(m)) \<longleftrightarrow> (w $< z $+ $#m | intify(w) = z $+ $#m)"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
142  | 
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
 | 
143  | 
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
 | 
144  | 
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
 | 
145  | 
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
 | 
146  | 
apply simp  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
147  | 
apply (erule natE)  | 
| 
 
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  | 
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
 | 
150  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
151  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
152  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
153  | 
lemma zadd_succ_lemma:  | 
| 61395 | 154  | 
"z \<in> int ==> (w $+ $# succ(m) $\<le> z) \<longleftrightarrow> (w $+ $#m $< z)"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
155  | 
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
 | 
156  | 
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
 | 
157  | 
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
 | 
158  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
159  | 
|
| 61395 | 160  | 
lemma zadd_succ_zle_iff: "(w $+ $# succ(m) $\<le> z) \<longleftrightarrow> (w $+ $#m $< z)"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
161  | 
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
 | 
162  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
163  | 
done  | 
| 
 
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  | 
(** Inequality reasoning **)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
166  | 
|
| 61395 | 167  | 
lemma zless_add1_iff_zle: "(w $< z $+ #1) \<longleftrightarrow> (w$\<le>z)"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
168  | 
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
 | 
169  | 
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
 | 
170  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
171  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
172  | 
|
| 61395 | 173  | 
lemma add1_zle_iff: "(w $+ #1 $\<le> z) \<longleftrightarrow> (w $< z)"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
174  | 
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
 | 
175  | 
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
 | 
176  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
177  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
178  | 
|
| 61395 | 179  | 
lemma add1_left_zle_iff: "(#1 $+ w $\<le> z) \<longleftrightarrow> (w $< z)"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
180  | 
apply (subst zadd_commute)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
181  | 
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
 | 
182  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
183  | 
|
| 
 
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  | 
(*** Monotonicity of Multiplication ***)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
186  | 
|
| 61395 | 187  | 
lemma zmult_mono_lemma: "k \<in> nat ==> i $\<le> j ==> i $* $#k $\<le> j $* $#k"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
188  | 
apply (induct_tac "k")  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
189  | 
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
 | 
190  | 
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
 | 
191  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
192  | 
|
| 61395 | 193  | 
lemma zmult_zle_mono1: "[| i $\<le> j; #0 $\<le> k |] ==> i$*k $\<le> j$*k"  | 
194  | 
apply (subgoal_tac "i $* intify (k) $\<le> j $* intify (k) ")  | 
|
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
195  | 
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
 | 
196  | 
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
 | 
197  | 
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
 | 
198  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
199  | 
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
 | 
200  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
201  | 
|
| 61395 | 202  | 
lemma zmult_zle_mono1_neg: "[| i $\<le> j; k $\<le> #0 |] ==> j$*k $\<le> i$*k"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
203  | 
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
 | 
204  | 
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
 | 
205  | 
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
 | 
206  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
207  | 
|
| 61395 | 208  | 
lemma zmult_zle_mono2: "[| i $\<le> j; #0 $\<le> k |] ==> k$*i $\<le> k$*j"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
209  | 
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
 | 
210  | 
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
 | 
211  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
212  | 
|
| 61395 | 213  | 
lemma zmult_zle_mono2_neg: "[| i $\<le> j; k $\<le> #0 |] ==> k$*j $\<le> k$*i"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
214  | 
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
 | 
215  | 
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
 | 
216  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
217  | 
|
| 61395 | 218  | 
(* $\<le> monotonicity, BOTH arguments*)  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
219  | 
lemma zmult_zle_mono:  | 
| 61395 | 220  | 
"[| i $\<le> j; k $\<le> l; #0 $\<le> j; #0 $\<le> k |] ==> i$*k $\<le> j$*l"  | 
| 
26056
 
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_mono1 [THEN zle_trans])  | 
| 
 
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  | 
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
 | 
224  | 
apply assumption  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
225  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
226  | 
|
| 
 
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  | 
(** 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
 | 
229  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
230  | 
lemma zmult_zless_mono2_lemma [rule_format]:  | 
| 46820 | 231  | 
"[| i$<j; k \<in> nat |] ==> 0<k \<longrightarrow> $#k $* i $< $#k $* j"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
232  | 
apply (induct_tac "k")  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
233  | 
prefer 2  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
234  | 
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
 | 
235  | 
apply (erule natE)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
236  | 
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
 | 
237  | 
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
 | 
238  | 
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
 | 
239  | 
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
 | 
240  | 
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
 | 
241  | 
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
 | 
242  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
243  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
244  | 
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
 | 
245  | 
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
 | 
246  | 
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
 | 
247  | 
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
 | 
248  | 
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
 | 
249  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
250  | 
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
 | 
251  | 
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
 | 
252  | 
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
 | 
253  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
254  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
255  | 
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
 | 
256  | 
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
 | 
257  | 
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
 | 
258  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
259  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
260  | 
(* < monotonicity, BOTH arguments*)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
261  | 
lemma zmult_zless_mono:  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
262  | 
"[| 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
 | 
263  | 
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
 | 
264  | 
apply assumption  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
265  | 
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
 | 
266  | 
apply assumption  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
267  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
268  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
269  | 
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
 | 
270  | 
apply (rule zminus_zless_zminus [THEN iffD1])  | 
| 46820 | 271  | 
apply (simp del: zmult_zminus_right  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
272  | 
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
 | 
273  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
274  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
275  | 
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
 | 
276  | 
apply (rule zminus_zless_zminus [THEN iffD1])  | 
| 46820 | 277  | 
apply (simp del: zmult_zminus  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
278  | 
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
 | 
279  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
280  | 
|
| 
 
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  | 
(** Products of zeroes **)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
283  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
284  | 
lemma zmult_eq_lemma:  | 
| 
46821
 
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
 
paulson 
parents: 
46820 
diff
changeset
 | 
285  | 
"[| m \<in> int; n \<in> int |] ==> (m = #0 | n = #0) \<longleftrightarrow> (m$*n = #0)"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
286  | 
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
 | 
287  | 
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
 | 
288  | 
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
 | 
289  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
290  | 
|
| 
46821
 
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
 
paulson 
parents: 
46820 
diff
changeset
 | 
291  | 
lemma zmult_eq_0_iff [iff]: "(m$*n = #0) \<longleftrightarrow> (intify(m) = #0 | intify(n) = #0)"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
292  | 
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
 | 
293  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
294  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
295  | 
|
| 46820 | 296  | 
(** Cancellation laws for k*m < k*n and m*k < n*k, also for @{text"\<le>"} and =,
 | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
297  | 
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
 | 
298  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
299  | 
lemma zmult_zless_lemma:  | 
| 46820 | 300  | 
"[| k \<in> int; m \<in> int; n \<in> int |]  | 
| 
46821
 
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
 
paulson 
parents: 
46820 
diff
changeset
 | 
301  | 
==> (m$*k $< n$*k) \<longleftrightarrow> ((#0 $< k & m$<n) | (k $< #0 & n$<m))"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
302  | 
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
 | 
303  | 
apply (auto simp add: neq_iff_zless zmult_zless_mono1 zmult_zless_mono1_neg)  | 
| 46820 | 304  | 
apply (auto simp add: not_zless_iff_zle  | 
305  | 
not_zle_iff_zless [THEN iff_sym, of "m$*k"]  | 
|
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
306  | 
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
 | 
307  | 
apply (auto elim: notE  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
308  | 
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
 | 
309  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
310  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
311  | 
lemma zmult_zless_cancel2:  | 
| 
46821
 
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
 
paulson 
parents: 
46820 
diff
changeset
 | 
312  | 
"(m$*k $< n$*k) \<longleftrightarrow> ((#0 $< k & m$<n) | (k $< #0 & n$<m))"  | 
| 46820 | 313  | 
apply (cut_tac k = "intify (k)" and m = "intify (m)" and n = "intify (n)"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
314  | 
in zmult_zless_lemma)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
315  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
316  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
317  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
318  | 
lemma zmult_zless_cancel1:  | 
| 
46821
 
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
 
paulson 
parents: 
46820 
diff
changeset
 | 
319  | 
"(k$*m $< k$*n) \<longleftrightarrow> ((#0 $< k & m$<n) | (k $< #0 & n$<m))"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
320  | 
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
 | 
321  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
322  | 
lemma zmult_zle_cancel2:  | 
| 61395 | 323  | 
"(m$*k $\<le> n$*k) \<longleftrightarrow> ((#0 $< k \<longrightarrow> m$\<le>n) & (k $< #0 \<longrightarrow> n$\<le>m))"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
324  | 
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
 | 
325  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
326  | 
lemma zmult_zle_cancel1:  | 
| 61395 | 327  | 
"(k$*m $\<le> k$*n) \<longleftrightarrow> ((#0 $< k \<longrightarrow> m$\<le>n) & (k $< #0 \<longrightarrow> n$\<le>m))"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
328  | 
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
 | 
329  | 
|
| 61395 | 330  | 
lemma int_eq_iff_zle: "[| m \<in> int; n \<in> int |] ==> m=n \<longleftrightarrow> (m $\<le> n & n $\<le> m)"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
331  | 
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
 | 
332  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
333  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
334  | 
lemma zmult_cancel2_lemma:  | 
| 
46821
 
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
 
paulson 
parents: 
46820 
diff
changeset
 | 
335  | 
"[| k \<in> int; m \<in> int; n \<in> int |] ==> (m$*k = n$*k) \<longleftrightarrow> (k=#0 | m=n)"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
336  | 
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
 | 
337  | 
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
 | 
338  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
339  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
340  | 
lemma zmult_cancel2 [simp]:  | 
| 
46821
 
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
 
paulson 
parents: 
46820 
diff
changeset
 | 
341  | 
"(m$*k = n$*k) \<longleftrightarrow> (intify(k) = #0 | intify(m) = intify(n))"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
342  | 
apply (rule iff_trans)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
343  | 
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
 | 
344  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
345  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
346  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
347  | 
lemma zmult_cancel1 [simp]:  | 
| 
46821
 
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
 
paulson 
parents: 
46820 
diff
changeset
 | 
348  | 
"(k$*m = k$*n) \<longleftrightarrow> (intify(k) = #0 | intify(m) = intify(n))"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
349  | 
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
 | 
350  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
351  | 
|
| 60770 | 352  | 
subsection\<open>Uniqueness and monotonicity of quotients and remainders\<close>  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
353  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
354  | 
lemma unique_quotient_lemma:  | 
| 61395 | 355  | 
"[| b$*q' $+ r' $\<le> b$*q $+ r; #0 $\<le> r'; #0 $< b; r $< b |]  | 
356  | 
==> q' $\<le> q"  | 
|
357  | 
apply (subgoal_tac "r' $+ b $* (q'$-q) $\<le> r")  | 
|
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
358  | 
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
 | 
359  | 
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
 | 
360  | 
prefer 2  | 
| 
 
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: 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
 | 
363  | 
apply (erule zle_zless_trans)  | 
| 46993 | 364  | 
apply simp  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
365  | 
apply (subgoal_tac "b $* q' $< b $* (#1 $+ q)")  | 
| 46820 | 366  | 
prefer 2  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
367  | 
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
 | 
368  | 
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
 | 
369  | 
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
 | 
370  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
371  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
372  | 
lemma unique_quotient_lemma_neg:  | 
| 61395 | 373  | 
"[| b$*q' $+ r' $\<le> b$*q $+ r; r $\<le> #0; b $< #0; b $< r' |]  | 
374  | 
==> q $\<le> q'"  | 
|
| 46820 | 375  | 
apply (rule_tac b = "$-b" and r = "$-r'" and r' = "$-r"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
376  | 
in unique_quotient_lemma)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
377  | 
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
 | 
378  | 
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
 | 
379  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
380  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
381  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
382  | 
lemma unique_quotient:  | 
| 46820 | 383  | 
"[| quorem (<a,b>, <q,r>); quorem (<a,b>, <q',r'>); b \<in> int; b \<noteq> #0;  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
384  | 
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
 | 
385  | 
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
 | 
386  | 
apply safe  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
387  | 
apply simp_all  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
388  | 
apply (blast intro: zle_anti_sym  | 
| 46820 | 389  | 
dest: zle_eq_refl [THEN unique_quotient_lemma]  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
390  | 
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
 | 
391  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
392  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
393  | 
lemma unique_remainder:  | 
| 46820 | 394  | 
"[| quorem (<a,b>, <q,r>); quorem (<a,b>, <q',r'>); b \<in> int; b \<noteq> #0;  | 
395  | 
q \<in> int; q' \<in> int;  | 
|
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
396  | 
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
 | 
397  | 
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
 | 
398  | 
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
 | 
399  | 
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
 | 
400  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
401  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
402  | 
|
| 60770 | 403  | 
subsection\<open>Correctness of posDivAlg,  | 
| 61798 | 404  | 
the Division Algorithm for \<open>a\<ge>0\<close> and \<open>b>0\<close>\<close>  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
405  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
406  | 
lemma adjust_eq [simp]:  | 
| 46820 | 407  | 
"adjust(b, <q,r>) = (let diff = r$-b in  | 
| 61395 | 408  | 
if #0 $\<le> diff then <#2$*q $+ #1,diff>  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
409  | 
else <#2$*q,r>)"  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
410  | 
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
 | 
411  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
412  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
413  | 
lemma posDivAlg_termination:  | 
| 46820 | 414  | 
"[| #0 $< b; ~ a $< b |]  | 
| 61395 | 415  | 
==> nat_of(a $- #2 $* b $+ #1) < nat_of(a $- b $+ #1)"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
416  | 
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
 | 
417  | 
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
 | 
418  | 
done  | 
| 
 
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  | 
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
 | 
421  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
422  | 
lemma posDivAlg_eqn:  | 
| 46820 | 423  | 
"[| #0 $< b; a \<in> int; b \<in> int |] ==>  | 
424  | 
posDivAlg(<a,b>) =  | 
|
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
425  | 
(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
 | 
426  | 
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
 | 
427  | 
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
 | 
428  | 
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
 | 
429  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
430  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
431  | 
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
 | 
432  | 
assumes prem:  | 
| 46820 | 433  | 
"!!a b. [| a \<in> int; b \<in> int;  | 
| 61395 | 434  | 
~ (a $< b | b $\<le> #0) \<longrightarrow> P(<a, #2 $* b>) |] ==> P(<a,b>)"  | 
| 46993 | 435  | 
shows "<u,v> \<in> int*int \<Longrightarrow> P(<u,v>)"  | 
436  | 
using wf_measure [where A = "int*int" and f = "%<a,b>.nat_of (a $- b $+ #1)"]  | 
|
437  | 
proof (induct "<u,v>" arbitrary: u v rule: wf_induct)  | 
|
438  | 
case (step x)  | 
|
439  | 
hence uv: "u \<in> int" "v \<in> int" by auto  | 
|
440  | 
thus ?case  | 
|
441  | 
apply (rule prem)  | 
|
442  | 
apply (rule impI)  | 
|
443  | 
apply (rule step)  | 
|
444  | 
apply (auto simp add: step uv not_zle_iff_zless posDivAlg_termination)  | 
|
445  | 
done  | 
|
446  | 
qed  | 
|
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
447  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
448  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
449  | 
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
 | 
450  | 
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
 | 
451  | 
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
 | 
452  | 
and ih: "!!a b. [| a \<in> int; b \<in> int;  | 
| 61395 | 453  | 
~ (a $< b | b $\<le> #0) \<longrightarrow> P(a, #2 $* b) |] ==> P(a,b)"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
454  | 
shows "P(u,v)"  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
455  | 
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
 | 
456  | 
apply simp  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
457  | 
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
 | 
458  | 
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
 | 
459  | 
apply (rule ih)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
460  | 
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
 | 
461  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
462  | 
|
| 46820 | 463  | 
(*FIXME: use intify in integ_of so that we always have @{term"integ_of w \<in> int"}.
 | 
464  | 
then this rewrite can work for all constants!!*)  | 
|
| 61395 | 465  | 
lemma intify_eq_0_iff_zle: "intify(m) = #0 \<longleftrightarrow> (m $\<le> #0 & #0 $\<le> m)"  | 
| 46993 | 466  | 
by (simp add: int_eq_iff_zle)  | 
| 
26056
 
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  | 
|
| 60770 | 469  | 
subsection\<open>Some convenient biconditionals for products of signs\<close>  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
470  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
471  | 
lemma zmult_pos: "[| #0 $< i; #0 $< j |] ==> #0 $< i $* j"  | 
| 46993 | 472  | 
by (drule zmult_zless_mono1, auto)  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
473  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
474  | 
lemma zmult_neg: "[| i $< #0; j $< #0 |] ==> #0 $< i $* j"  | 
| 46993 | 475  | 
by (drule zmult_zless_mono1_neg, auto)  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
476  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
477  | 
lemma zmult_pos_neg: "[| #0 $< i; j $< #0 |] ==> i $* j $< #0"  | 
| 46993 | 478  | 
by (drule zmult_zless_mono1_neg, auto)  | 
479  | 
||
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
480  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
481  | 
(** Inequality reasoning **)  | 
| 
 
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  | 
lemma int_0_less_lemma:  | 
| 46820 | 484  | 
"[| x \<in> int; y \<in> int |]  | 
| 
46821
 
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
 
paulson 
parents: 
46820 
diff
changeset
 | 
485  | 
==> (#0 $< x $* y) \<longleftrightarrow> (#0 $< x & #0 $< y | x $< #0 & y $< #0)"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
486  | 
apply (auto simp add: zle_def not_zless_iff_zle zmult_pos zmult_neg)  | 
| 46820 | 487  | 
apply (rule ccontr)  | 
488  | 
apply (rule_tac [2] ccontr)  | 
|
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
489  | 
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
 | 
490  | 
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
 | 
491  | 
apply (erule_tac [2] P = "#0$< x$* y" in rev_mp)  | 
| 46820 | 492  | 
apply (drule zmult_pos_neg, assumption)  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
493  | 
prefer 2  | 
| 46820 | 494  | 
apply (drule zmult_pos_neg, assumption)  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
495  | 
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
 | 
496  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
497  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
498  | 
lemma int_0_less_mult_iff:  | 
| 
46821
 
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
 
paulson 
parents: 
46820 
diff
changeset
 | 
499  | 
"(#0 $< x $* y) \<longleftrightarrow> (#0 $< x & #0 $< y | x $< #0 & y $< #0)"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
500  | 
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
 | 
501  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
502  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
503  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
504  | 
lemma int_0_le_lemma:  | 
| 46820 | 505  | 
"[| x \<in> int; y \<in> int |]  | 
| 61395 | 506  | 
==> (#0 $\<le> x $* y) \<longleftrightarrow> (#0 $\<le> x & #0 $\<le> y | x $\<le> #0 & y $\<le> #0)"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
507  | 
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
 | 
508  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
509  | 
lemma int_0_le_mult_iff:  | 
| 61395 | 510  | 
"(#0 $\<le> x $* y) \<longleftrightarrow> ((#0 $\<le> x & #0 $\<le> y) | (x $\<le> #0 & y $\<le> #0))"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
511  | 
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
 | 
512  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
513  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
514  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
515  | 
lemma zmult_less_0_iff:  | 
| 
46821
 
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
 
paulson 
parents: 
46820 
diff
changeset
 | 
516  | 
"(x $* y $< #0) \<longleftrightarrow> (#0 $< x & y $< #0 | x $< #0 & #0 $< y)"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
517  | 
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
 | 
518  | 
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
 | 
519  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
520  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
521  | 
lemma zmult_le_0_iff:  | 
| 61395 | 522  | 
"(x $* y $\<le> #0) \<longleftrightarrow> (#0 $\<le> x & y $\<le> #0 | x $\<le> #0 & #0 $\<le> y)"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
523  | 
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
 | 
524  | 
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
 | 
525  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
526  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
527  | 
(*Typechecking for posDivAlg*)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
528  | 
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
 | 
529  | 
"[| 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
 | 
530  | 
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
 | 
531  | 
apply assumption+  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
532  | 
apply (case_tac "#0 $< ba")  | 
| 46820 | 533  | 
apply (simp add: posDivAlg_eqn adjust_def integ_of_type  | 
| 63648 | 534  | 
split: split_if_asm)  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
535  | 
apply clarify  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
536  | 
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
 | 
537  | 
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
 | 
538  | 
apply (subst posDivAlg_unfold)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
539  | 
apply simp  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
540  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
541  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
542  | 
(*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
 | 
543  | 
lemma posDivAlg_correct [rule_format]:  | 
| 46820 | 544  | 
"[| a \<in> int; b \<in> int |]  | 
| 61395 | 545  | 
==> #0 $\<le> a \<longrightarrow> #0 $< b \<longrightarrow> quorem (<a,b>, posDivAlg(<a,b>))"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
546  | 
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
 | 
547  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
548  | 
apply (simp_all add: quorem_def)  | 
| 60770 | 549  | 
txt\<open>base case: a<b\<close>  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
550  | 
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
 | 
551  | 
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
 | 
552  | 
apply (simp add: int_0_less_mult_iff)  | 
| 60770 | 553  | 
txt\<open>main argument\<close>  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
554  | 
apply (subst posDivAlg_eqn)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
555  | 
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
 | 
556  | 
apply (erule splitE)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
557  | 
apply (rule posDivAlg_type)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
558  | 
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
 | 
559  | 
apply (auto simp add: zadd_zmult_distrib2 Let_def)  | 
| 60770 | 560  | 
txt\<open>now just linear arithmetic\<close>  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
561  | 
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
 | 
562  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
563  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
564  | 
|
| 60770 | 565  | 
subsection\<open>Correctness of negDivAlg, the division algorithm for a<0 and b>0\<close>  | 
| 
26056
 
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  | 
lemma negDivAlg_termination:  | 
| 46820 | 568  | 
"[| #0 $< b; a $+ b $< #0 |]  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
569  | 
==> 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
 | 
570  | 
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
 | 
571  | 
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
 | 
572  | 
zless_zminus)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
573  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
574  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
575  | 
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
 | 
576  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
577  | 
lemma negDivAlg_eqn:  | 
| 46820 | 578  | 
"[| #0 $< b; a \<in> int; b \<in> int |] ==>  | 
579  | 
negDivAlg(<a,b>) =  | 
|
| 61395 | 580  | 
(if #0 $\<le> a$+b 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
 | 
581  | 
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
 | 
582  | 
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
 | 
583  | 
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
 | 
584  | 
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
 | 
585  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
586  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
587  | 
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
 | 
588  | 
assumes prem:  | 
| 46820 | 589  | 
"!!a b. [| a \<in> int; b \<in> int;  | 
| 61395 | 590  | 
~ (#0 $\<le> a $+ b | b $\<le> #0) \<longrightarrow> P(<a, #2 $* b>) |]  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
591  | 
==> P(<a,b>)"  | 
| 46993 | 592  | 
shows "<u,v> \<in> int*int \<Longrightarrow> P(<u,v>)"  | 
593  | 
using wf_measure [where A = "int*int" and f = "%<a,b>.nat_of ($- a $- b)"]  | 
|
594  | 
proof (induct "<u,v>" arbitrary: u v rule: wf_induct)  | 
|
595  | 
case (step x)  | 
|
596  | 
hence uv: "u \<in> int" "v \<in> int" by auto  | 
|
597  | 
thus ?case  | 
|
598  | 
apply (rule prem)  | 
|
599  | 
apply (rule impI)  | 
|
600  | 
apply (rule step)  | 
|
601  | 
apply (auto simp add: step uv not_zle_iff_zless negDivAlg_termination)  | 
|
602  | 
done  | 
|
603  | 
qed  | 
|
| 
26056
 
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"  | 
| 46820 | 608  | 
and ih: "!!a b. [| a \<in> int; b \<in> int;  | 
| 61395 | 609  | 
~ (#0 $\<le> a $+ b | b $\<le> #0) \<longrightarrow> P(a, #2 $* b) |]  | 
| 
26056
 
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")  | 
| 46820 | 627  | 
apply (simp add: negDivAlg_eqn adjust_def integ_of_type  | 
| 63648 | 628  | 
split: split_if_asm)  | 
| 
26056
 
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]:  | 
| 46820 | 640  | 
"[| a \<in> int; b \<in> int |]  | 
641  | 
==> a $< #0 \<longrightarrow> #0 $< b \<longrightarrow> quorem (<a,b>, negDivAlg(<a,b>))"  | 
|
| 
26056
 
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)  | 
| 61395 | 645  | 
   txt\<open>base case: @{term "0$\<le>a$+b"}\<close>
 | 
| 
26056
 
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)  | 
| 60770 | 649  | 
txt\<open>main argument\<close>  | 
| 
26056
 
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)  | 
| 60770 | 656  | 
txt\<open>now just linear arithmetic\<close>  | 
| 
26056
 
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  | 
|
| 60770 | 661  | 
subsection\<open>Existence shown by proving the division algorithm to be correct\<close>  | 
| 
26056
 
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.*)  | 
| 61395 | 679  | 
lemma linear_arith_lemma: "~ (#0 $\<le> #-1 $+ b) ==> (b $\<le> #0)"  | 
| 
26056
 
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:  | 
| 46820 | 701  | 
"[|quorem (<$-a,$-b>, qr); a \<in> int; b \<in> int; qr \<in> int * int|]  | 
| 
26056
 
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)  | 
| 60770 | 705  | 
txt\<open>linear arithmetic from here on\<close>  | 
| 
26056
 
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  | 
| 46820 | 717  | 
posDivAlg_type negDivAlg_type)  | 
| 
26056
 
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)  | 
| 60770 | 719  | 
txt\<open>linear arithmetic from here on\<close>  | 
| 
26056
 
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"  | 
| 46993 | 732  | 
by (simp add: zdiv_def)  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
733  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
734  | 
lemma zdiv_intify2 [simp]: "x zdiv intify(y) = x zdiv y"  | 
| 46993 | 735  | 
by (simp add: zdiv_def)  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
736  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
737  | 
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
 | 
738  | 
apply (unfold zdiv_def)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
739  | 
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
 | 
740  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
741  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
742  | 
lemma zmod_intify1 [simp]: "intify(x) zmod y = x zmod y"  | 
| 46993 | 743  | 
by (simp add: zmod_def)  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
744  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
745  | 
lemma zmod_intify2 [simp]: "x zmod intify(y) = x zmod y"  | 
| 46993 | 746  | 
by (simp add: zmod_def)  | 
| 
26056
 
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_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
 | 
749  | 
apply (unfold zmod_def)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
750  | 
apply (rule snd_type)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
751  | 
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
 | 
752  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
753  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
754  | 
|
| 46820 | 755  | 
(** Arbitrary definitions for division by zero. Useful to simplify  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
756  | 
certain equations **)  | 
| 
 
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  | 
lemma DIVISION_BY_ZERO_ZDIV: "a zdiv #0 = #0"  | 
| 46993 | 759  | 
by (simp add: zdiv_def divAlg_def posDivAlg_zero_divisor)  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
760  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
761  | 
lemma DIVISION_BY_ZERO_ZMOD: "a zmod #0 = intify(a)"  | 
| 46993 | 762  | 
by (simp add: zmod_def divAlg_def posDivAlg_zero_divisor)  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
763  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
764  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
765  | 
(** 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
 | 
766  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
767  | 
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
 | 
768  | 
"[| 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
 | 
769  | 
apply (case_tac "b = #0")  | 
| 46820 | 770  | 
apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD)  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
771  | 
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
 | 
772  | 
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
 | 
773  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
774  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
775  | 
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
 | 
776  | 
apply (rule trans)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
777  | 
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
 | 
778  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
779  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
780  | 
|
| 61395 | 781  | 
lemma pos_mod: "#0 $< b ==> #0 $\<le> a zmod b & a zmod b $< b"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
782  | 
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
 | 
783  | 
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
 | 
784  | 
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
 | 
785  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
786  | 
|
| 45602 | 787  | 
lemmas pos_mod_sign = pos_mod [THEN conjunct1]  | 
788  | 
and pos_mod_bound = pos_mod [THEN conjunct2]  | 
|
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
789  | 
|
| 61395 | 790  | 
lemma neg_mod: "b $< #0 ==> a zmod b $\<le> #0 & b $< a zmod b"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
791  | 
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
 | 
792  | 
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
 | 
793  | 
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
 | 
794  | 
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
 | 
795  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
796  | 
|
| 45602 | 797  | 
lemmas neg_mod_sign = neg_mod [THEN conjunct1]  | 
798  | 
and neg_mod_bound = neg_mod [THEN conjunct2]  | 
|
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
799  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
800  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
801  | 
(** 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
 | 
802  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
803  | 
lemma quorem_div_mod:  | 
| 46820 | 804  | 
"[|b \<noteq> #0; a \<in> int; b \<in> int |]  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
805  | 
==> 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
 | 
806  | 
apply (cut_tac a = "a" and b = "b" in zmod_zdiv_equality)  | 
| 46820 | 807  | 
apply (auto simp add: quorem_def neq_iff_zless pos_mod_sign pos_mod_bound  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
808  | 
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
 | 
809  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
810  | 
|
| 46820 | 811  | 
(*Surely quorem(<a,b>,<q,r>) implies @{term"a \<in> int"}, but it doesn't matter*)
 | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
812  | 
lemma quorem_div:  | 
| 46820 | 813  | 
"[| quorem(<a,b>,<q,r>); b \<noteq> #0; a \<in> int; b \<in> int; q \<in> int |]  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
814  | 
==> a zdiv b = q"  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
815  | 
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
 | 
816  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
817  | 
lemma quorem_mod:  | 
| 46820 | 818  | 
"[| quorem(<a,b>,<q,r>); b \<noteq> #0; a \<in> int; b \<in> int; q \<in> int; r \<in> int |]  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
819  | 
==> a zmod b = r"  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
820  | 
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
 | 
821  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
822  | 
lemma zdiv_pos_pos_trivial_raw:  | 
| 61395 | 823  | 
"[| a \<in> int; b \<in> int; #0 $\<le> a; a $< b |] ==> a zdiv b = #0"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
824  | 
apply (rule quorem_div)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
825  | 
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
 | 
826  | 
(*linear arithmetic*)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
827  | 
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
 | 
828  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
829  | 
|
| 61395 | 830  | 
lemma zdiv_pos_pos_trivial: "[| #0 $\<le> a; a $< b |] ==> a zdiv b = #0"  | 
| 46820 | 831  | 
apply (cut_tac a = "intify (a)" and b = "intify (b)"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
832  | 
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
 | 
833  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
834  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
835  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
836  | 
lemma zdiv_neg_neg_trivial_raw:  | 
| 61395 | 837  | 
"[| a \<in> int; b \<in> int; a $\<le> #0; b $< a |] ==> a zdiv b = #0"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
838  | 
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
 | 
839  | 
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
 | 
840  | 
(*linear arithmetic*)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
841  | 
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
 | 
842  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
843  | 
|
| 61395 | 844  | 
lemma zdiv_neg_neg_trivial: "[| a $\<le> #0; b $< a |] ==> a zdiv b = #0"  | 
| 46820 | 845  | 
apply (cut_tac a = "intify (a)" and b = "intify (b)"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
846  | 
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
 | 
847  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
848  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
849  | 
|
| 61395 | 850  | 
lemma zadd_le_0_lemma: "[| a$+b $\<le> #0; #0 $< a; #0 $< b |] ==> False"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
851  | 
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
 | 
852  | 
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
 | 
853  | 
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
 | 
854  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
855  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
856  | 
lemma zdiv_pos_neg_trivial_raw:  | 
| 61395 | 857  | 
"[| a \<in> int; b \<in> int; #0 $< a; a$+b $\<le> #0 |] ==> a zdiv b = #-1"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
858  | 
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
 | 
859  | 
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
 | 
860  | 
(*linear arithmetic*)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
861  | 
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
 | 
862  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
863  | 
|
| 61395 | 864  | 
lemma zdiv_pos_neg_trivial: "[| #0 $< a; a$+b $\<le> #0 |] ==> a zdiv b = #-1"  | 
| 46820 | 865  | 
apply (cut_tac a = "intify (a)" and b = "intify (b)"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
866  | 
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
 | 
867  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
868  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
869  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
870  | 
(*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
 | 
871  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
872  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
873  | 
lemma zmod_pos_pos_trivial_raw:  | 
| 61395 | 874  | 
"[| a \<in> int; b \<in> int; #0 $\<le> a; a $< b |] ==> a zmod b = a"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
875  | 
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
 | 
876  | 
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
 | 
877  | 
(*linear arithmetic*)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
878  | 
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
 | 
879  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
880  | 
|
| 61395 | 881  | 
lemma zmod_pos_pos_trivial: "[| #0 $\<le> a; a $< b |] ==> a zmod b = intify(a)"  | 
| 46820 | 882  | 
apply (cut_tac a = "intify (a)" and b = "intify (b)"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
883  | 
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
 | 
884  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
885  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
886  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
887  | 
lemma zmod_neg_neg_trivial_raw:  | 
| 61395 | 888  | 
"[| a \<in> int; b \<in> int; a $\<le> #0; b $< a |] ==> a zmod b = a"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
889  | 
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
 | 
890  | 
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
 | 
891  | 
(*linear arithmetic*)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
892  | 
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
 | 
893  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
894  | 
|
| 61395 | 895  | 
lemma zmod_neg_neg_trivial: "[| a $\<le> #0; b $< a |] ==> a zmod b = intify(a)"  | 
| 46820 | 896  | 
apply (cut_tac a = "intify (a)" and b = "intify (b)"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
897  | 
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
 | 
898  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
899  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
900  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
901  | 
lemma zmod_pos_neg_trivial_raw:  | 
| 61395 | 902  | 
"[| a \<in> int; b \<in> int; #0 $< a; a$+b $\<le> #0 |] ==> a zmod b = a$+b"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
903  | 
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
 | 
904  | 
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
 | 
905  | 
(*linear arithmetic*)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
906  | 
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
 | 
907  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
908  | 
|
| 61395 | 909  | 
lemma zmod_pos_neg_trivial: "[| #0 $< a; a$+b $\<le> #0 |] ==> a zmod b = a$+b"  | 
| 46820 | 910  | 
apply (cut_tac a = "intify (a)" and b = "intify (b)"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
911  | 
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
 | 
912  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
913  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
914  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
915  | 
(*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
 | 
916  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
917  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
918  | 
(*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
 | 
919  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
920  | 
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
 | 
921  | 
"[|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
 | 
922  | 
apply (case_tac "b = #0")  | 
| 46820 | 923  | 
apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD)  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
924  | 
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
 | 
925  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
926  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
927  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
928  | 
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
 | 
929  | 
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
 | 
930  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
931  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
932  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
933  | 
(*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
 | 
934  | 
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
 | 
935  | 
"[|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
 | 
936  | 
apply (case_tac "b = #0")  | 
| 46820 | 937  | 
apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD)  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
938  | 
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
 | 
939  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
940  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
941  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
942  | 
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
 | 
943  | 
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
 | 
944  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
945  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
946  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
947  | 
|
| 60770 | 948  | 
subsection\<open>division of a number by itself\<close>  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
949  | 
|
| 61395 | 950  | 
lemma self_quotient_aux1: "[| #0 $< a; a = r $+ a$*q; r $< a |] ==> #1 $\<le> q"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
951  | 
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
 | 
952  | 
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
 | 
953  | 
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
 | 
954  | 
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
 | 
955  | 
(*linear arithmetic...*)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
956  | 
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
 | 
957  | 
apply (drule sym)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
958  | 
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
 | 
959  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
960  | 
|
| 61395 | 961  | 
lemma self_quotient_aux2: "[| #0 $< a; a = r $+ a$*q; #0 $\<le> r |] ==> q $\<le> #1"  | 
962  | 
apply (subgoal_tac "#0 $\<le> a$* (#1$-q)")  | 
|
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
963  | 
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
 | 
964  | 
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
 | 
965  | 
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
 | 
966  | 
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
 | 
967  | 
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
 | 
968  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
969  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
970  | 
lemma self_quotient:  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
971  | 
"[| 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
 | 
972  | 
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
 | 
973  | 
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
 | 
974  | 
apply safe  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
975  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
976  | 
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
 | 
977  | 
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
 | 
978  | 
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
 | 
979  | 
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
 | 
980  | 
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
 | 
981  | 
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
 | 
982  | 
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
 | 
983  | 
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
 | 
984  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
985  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
986  | 
lemma self_remainder:  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
987  | 
"[|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
 | 
988  | 
apply (frule self_quotient)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
989  | 
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
 | 
990  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
991  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
992  | 
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
 | 
993  | 
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
 | 
994  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
995  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
996  | 
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
 | 
997  | 
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
 | 
998  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
999  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1000  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1001  | 
(*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
 | 
1002  | 
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
 | 
1003  | 
apply (case_tac "a = #0")  | 
| 46820 | 1004  | 
apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD)  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1005  | 
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
 | 
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  | 
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
 | 
1009  | 
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
 | 
1010  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1011  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1012  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1013  | 
|
| 60770 | 1014  | 
subsection\<open>Computation of division and remainder\<close>  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1015  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1016  | 
lemma zdiv_zero [simp]: "#0 zdiv b = #0"  | 
| 46993 | 1017  | 
by (simp add: zdiv_def divAlg_def)  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1018  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1019  | 
lemma zdiv_eq_minus1: "#0 $< b ==> #-1 zdiv b = #-1"  | 
| 46993 | 1020  | 
by (simp (no_asm_simp) add: zdiv_def divAlg_def)  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1021  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1022  | 
lemma zmod_zero [simp]: "#0 zmod b = #0"  | 
| 46993 | 1023  | 
by (simp add: zmod_def divAlg_def)  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1024  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1025  | 
lemma zdiv_minus1: "#0 $< b ==> #-1 zdiv b = #-1"  | 
| 46993 | 1026  | 
by (simp add: zdiv_def divAlg_def)  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1027  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1028  | 
lemma zmod_minus1: "#0 $< b ==> #-1 zmod b = b $- #1"  | 
| 46993 | 1029  | 
by (simp add: zmod_def divAlg_def)  | 
| 
26056
 
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  | 
(** a positive, b positive **)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1032  | 
|
| 61395 | 1033  | 
lemma zdiv_pos_pos: "[| #0 $< a; #0 $\<le> b |]  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1034  | 
==> 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
 | 
1035  | 
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
 | 
1036  | 
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
 | 
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_pos_pos:  | 
| 61395 | 1040  | 
"[| #0 $< a; #0 $\<le> b |]  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1041  | 
==> 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
 | 
1042  | 
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
 | 
1043  | 
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
 | 
1044  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1045  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1046  | 
(** a negative, b positive **)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1047  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1048  | 
lemma zdiv_neg_pos:  | 
| 46820 | 1049  | 
"[| a $< #0; #0 $< b |]  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1050  | 
==> 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
 | 
1051  | 
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
 | 
1052  | 
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
 | 
1053  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1054  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1055  | 
lemma zmod_neg_pos:  | 
| 46820 | 1056  | 
"[| a $< #0; #0 $< b |]  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1057  | 
==> 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
 | 
1058  | 
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
 | 
1059  | 
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
 | 
1060  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1061  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1062  | 
(** a positive, b negative **)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1063  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1064  | 
lemma zdiv_pos_neg:  | 
| 46820 | 1065  | 
"[| #0 $< a; b $< #0 |]  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1066  | 
==> 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
 | 
1067  | 
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
 | 
1068  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1069  | 
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
 | 
1070  | 
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
 | 
1071  | 
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
 | 
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  | 
lemma zmod_pos_neg:  | 
| 46820 | 1075  | 
"[| #0 $< a; b $< #0 |]  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1076  | 
==> 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
 | 
1077  | 
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
 | 
1078  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1079  | 
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
 | 
1080  | 
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
 | 
1081  | 
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
 | 
1082  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1083  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1084  | 
(** a negative, b negative **)  | 
| 
 
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 zdiv_neg_neg:  | 
| 61395 | 1087  | 
"[| a $< #0; b $\<le> #0 |]  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1088  | 
==> 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
 | 
1089  | 
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
 | 
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  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1093  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1094  | 
lemma zmod_neg_neg:  | 
| 61395 | 1095  | 
"[| a $< #0; b $\<le> #0 |]  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1096  | 
==> 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
 | 
1097  | 
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
 | 
1098  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1099  | 
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
 | 
1100  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1101  | 
|
| 45602 | 1102  | 
declare zdiv_pos_pos [of "integ_of (v)" "integ_of (w)", simp] for v w  | 
1103  | 
declare zdiv_neg_pos [of "integ_of (v)" "integ_of (w)", simp] for v w  | 
|
1104  | 
declare zdiv_pos_neg [of "integ_of (v)" "integ_of (w)", simp] for v w  | 
|
1105  | 
declare zdiv_neg_neg [of "integ_of (v)" "integ_of (w)", simp] for v w  | 
|
1106  | 
declare zmod_pos_pos [of "integ_of (v)" "integ_of (w)", simp] for v w  | 
|
1107  | 
declare zmod_neg_pos [of "integ_of (v)" "integ_of (w)", simp] for v w  | 
|
1108  | 
declare zmod_pos_neg [of "integ_of (v)" "integ_of (w)", simp] for v w  | 
|
1109  | 
declare zmod_neg_neg [of "integ_of (v)" "integ_of (w)", simp] for v w  | 
|
1110  | 
declare posDivAlg_eqn [of concl: "integ_of (v)" "integ_of (w)", simp] for v w  | 
|
1111  | 
declare negDivAlg_eqn [of concl: "integ_of (v)" "integ_of (w)", simp] for v w  | 
|
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1112  | 
|
| 
 
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  | 
(** Special-case simplification **)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1115  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1116  | 
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
 | 
1117  | 
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
 | 
1118  | 
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
 | 
1119  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1120  | 
(*arithmetic*)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1121  | 
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
 | 
1122  | 
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
 | 
1123  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1124  | 
done  | 
| 
 
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  | 
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
 | 
1127  | 
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
 | 
1128  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1129  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1130  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1131  | 
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
 | 
1132  | 
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
 | 
1133  | 
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
 | 
1134  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1135  | 
(*arithmetic*)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1136  | 
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
 | 
1137  | 
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
 | 
1138  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1139  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1140  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1141  | 
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
 | 
1142  | 
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
 | 
1143  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1144  | 
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
 | 
1145  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1146  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1147  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1148  | 
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
 | 
1149  | 
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
 | 
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  | 
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
 | 
1153  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1154  | 
|
| 60770 | 1155  | 
subsection\<open>Monotonicity in the first argument (divisor)\<close>  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1156  | 
|
| 61395 | 1157  | 
lemma zdiv_mono1: "[| a $\<le> a'; #0 $< b |] ==> a zdiv b $\<le> a' zdiv b"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1158  | 
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
 | 
1159  | 
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
 | 
1160  | 
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
 | 
1161  | 
apply (erule subst)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1162  | 
apply (erule subst)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1163  | 
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
 | 
1164  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1165  | 
|
| 61395 | 1166  | 
lemma zdiv_mono1_neg: "[| a $\<le> a'; b $< #0 |] ==> a' zdiv b $\<le> a zdiv b"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1167  | 
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
 | 
1168  | 
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
 | 
1169  | 
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
 | 
1170  | 
apply (erule subst)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1171  | 
apply (erule subst)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1172  | 
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
 | 
1173  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1174  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1175  | 
|
| 60770 | 1176  | 
subsection\<open>Monotonicity in the second argument (dividend)\<close>  | 
| 
26056
 
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 q_pos_lemma:  | 
| 61395 | 1179  | 
"[| #0 $\<le> b'$*q' $+ r'; r' $< b'; #0 $< b' |] ==> #0 $\<le> q'"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1180  | 
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
 | 
1181  | 
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
 | 
1182  | 
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
 | 
1183  | 
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
 | 
1184  | 
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
 | 
1185  | 
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
 | 
1186  | 
done  | 
| 
 
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  | 
lemma zdiv_mono2_lemma:  | 
| 61395 | 1189  | 
"[| b$*q $+ r = b'$*q' $+ r'; #0 $\<le> b'$*q' $+ r';  | 
1190  | 
r' $< b'; #0 $\<le> r; #0 $< b'; b' $\<le> b |]  | 
|
1191  | 
==> q $\<le> q'"  | 
|
| 46820 | 1192  | 
apply (frule q_pos_lemma, assumption+)  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1193  | 
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
 | 
1194  | 
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
 | 
1195  | 
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
 | 
1196  | 
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
 | 
1197  | 
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
 | 
1198  | 
apply (simp (no_asm_simp) add: zadd_zmult_distrib2)  | 
| 61395 | 1199  | 
apply (subst zadd_commute [of "b $* q'"], rule zadd_zless_mono)  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1200  | 
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
 | 
1201  | 
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
 | 
1202  | 
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
 | 
1203  | 
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
 | 
1204  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1205  | 
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
 | 
1206  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1207  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1208  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1209  | 
lemma zdiv_mono2_raw:  | 
| 61395 | 1210  | 
"[| #0 $\<le> a; #0 $< b'; b' $\<le> b; a \<in> int |]  | 
1211  | 
==> a zdiv b $\<le> a zdiv b'"  | 
|
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1212  | 
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
 | 
1213  | 
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
 | 
1214  | 
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
 | 
1215  | 
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
 | 
1216  | 
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
 | 
1217  | 
apply (erule subst)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1218  | 
apply (erule subst)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1219  | 
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
 | 
1220  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1221  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1222  | 
lemma zdiv_mono2:  | 
| 61395 | 1223  | 
"[| #0 $\<le> a; #0 $< b'; b' $\<le> b |]  | 
1224  | 
==> a zdiv b $\<le> a zdiv b'"  | 
|
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1225  | 
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
 | 
1226  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1227  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1228  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1229  | 
lemma q_neg_lemma:  | 
| 61395 | 1230  | 
"[| b'$*q' $+ r' $< #0; #0 $\<le> r'; #0 $< b' |] ==> q' $< #0"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1231  | 
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
 | 
1232  | 
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
 | 
1233  | 
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
 | 
1234  | 
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
 | 
1235  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1236  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1237  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1238  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1239  | 
lemma zdiv_mono2_neg_lemma:  | 
| 46820 | 1240  | 
"[| b$*q $+ r = b'$*q' $+ r'; b'$*q' $+ r' $< #0;  | 
| 61395 | 1241  | 
r $< b; #0 $\<le> r'; #0 $< b'; b' $\<le> b |]  | 
1242  | 
==> q' $\<le> q"  | 
|
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1243  | 
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
 | 
1244  | 
prefer 2 apply (blast dest: zless_zle_trans)  | 
| 46820 | 1245  | 
apply (frule q_neg_lemma, assumption+)  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1246  | 
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
 | 
1247  | 
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
 | 
1248  | 
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
 | 
1249  | 
apply (simp (no_asm_simp) add: zadd_zmult_distrib2)  | 
| 61395 | 1250  | 
apply (subgoal_tac "b$*q' $\<le> b'$*q'")  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1251  | 
prefer 2  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1252  | 
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
 | 
1253  | 
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
 | 
1254  | 
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
 | 
1255  | 
prefer 2  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1256  | 
apply (erule ssubst)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1257  | 
apply simp  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1258  | 
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
 | 
1259  | 
apply (assumption)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1260  | 
apply simp  | 
| 
 
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_use) add: zadd_commute)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1262  | 
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
 | 
1263  | 
prefer 2 apply (assumption)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1264  | 
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
 | 
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  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1267  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1268  | 
lemma zdiv_mono2_neg_raw:  | 
| 61395 | 1269  | 
"[| a $< #0; #0 $< b'; b' $\<le> b; a \<in> int |]  | 
1270  | 
==> a zdiv b' $\<le> a zdiv b"  | 
|
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1271  | 
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
 | 
1272  | 
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
 | 
1273  | 
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
 | 
1274  | 
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
 | 
1275  | 
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
 | 
1276  | 
apply (erule subst)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1277  | 
apply (erule subst)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1278  | 
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
 | 
1279  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1280  | 
|
| 61395 | 1281  | 
lemma zdiv_mono2_neg: "[| a $< #0; #0 $< b'; b' $\<le> b |]  | 
1282  | 
==> a zdiv b' $\<le> a zdiv b"  | 
|
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1283  | 
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
 | 
1284  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1285  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1286  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1287  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1288  | 
|
| 60770 | 1289  | 
subsection\<open>More algebraic laws for zdiv and zmod\<close>  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1290  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1291  | 
(** 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
 | 
1292  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1293  | 
lemma zmult1_lemma:  | 
| 46820 | 1294  | 
"[| quorem(<b,c>, <q,r>); c \<in> int; c \<noteq> #0 |]  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1295  | 
==> 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
 | 
1296  | 
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
 | 
1297  | 
pos_mod_sign pos_mod_bound neg_mod_sign neg_mod_bound)  | 
| 46820 | 1298  | 
apply (auto intro: raw_zmod_zdiv_equality)  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1299  | 
done  | 
| 
 
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  | 
lemma zdiv_zmult1_eq_raw:  | 
| 46820 | 1302  | 
"[|b \<in> int; c \<in> int|]  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1303  | 
==> (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
 | 
1304  | 
apply (case_tac "c = #0")  | 
| 46820 | 1305  | 
apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD)  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1306  | 
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
 | 
1307  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1308  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1309  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1310  | 
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
 | 
1311  | 
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
 | 
1312  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1313  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1314  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1315  | 
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
 | 
1316  | 
"[|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
 | 
1317  | 
apply (case_tac "c = #0")  | 
| 46820 | 1318  | 
apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD)  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1319  | 
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
 | 
1320  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1321  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1322  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1323  | 
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
 | 
1324  | 
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
 | 
1325  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1326  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1327  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1328  | 
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
 | 
1329  | 
apply (rule trans)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1330  | 
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
 | 
1331  | 
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
 | 
1332  | 
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
 | 
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_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
 | 
1336  | 
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
 | 
1337  | 
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
 | 
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 zdiv_zmult_self1 [simp]: "intify(b) \<noteq> #0 ==> (a$*b) zdiv b = intify(a)"  | 
| 46993 | 1341  | 
by (simp add: zdiv_zmult1_eq)  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1342  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1343  | 
lemma zdiv_zmult_self2 [simp]: "intify(b) \<noteq> #0 ==> (b$*a) zdiv b = intify(a)"  | 
| 46993 | 1344  | 
by (simp add: zmult_commute)  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1345  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1346  | 
lemma zmod_zmult_self1 [simp]: "(a$*b) zmod b = #0"  | 
| 46993 | 1347  | 
by (simp add: zmod_zmult1_eq)  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1348  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1349  | 
lemma zmod_zmult_self2 [simp]: "(b$*a) zmod b = #0"  | 
| 46993 | 1350  | 
by (simp add: zmult_commute zmod_zmult1_eq)  | 
| 
26056
 
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  | 
|
| 46820 | 1353  | 
(** proving (a$+b) zdiv c =  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1354  | 
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
 | 
1355  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1356  | 
lemma zadd1_lemma:  | 
| 46820 | 1357  | 
"[| quorem(<a,c>, <aq,ar>); quorem(<b,c>, <bq,br>);  | 
1358  | 
c \<in> int; c \<noteq> #0 |]  | 
|
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1359  | 
==> 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
 | 
1360  | 
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
 | 
1361  | 
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
 | 
1362  | 
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
 | 
1363  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1364  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1365  | 
(*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
 | 
1366  | 
lemma zdiv_zadd1_eq_raw:  | 
| 46820 | 1367  | 
"[|a \<in> int; b \<in> int; c \<in> int|] ==>  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1368  | 
(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
 | 
1369  | 
apply (case_tac "c = #0")  | 
| 46820 | 1370  | 
apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD)  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1371  | 
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
 | 
1372  | 
THEN quorem_div])  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1373  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1374  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1375  | 
lemma zdiv_zadd1_eq:  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1376  | 
"(a$+b) zdiv c = a zdiv c $+ b zdiv c $+ ((a zmod c $+ b zmod c) zdiv c)"  | 
| 46820 | 1377  | 
apply (cut_tac a = "intify (a)" and b = "intify (b)" and c = "intify (c)"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1378  | 
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
 | 
1379  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1380  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1381  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1382  | 
lemma zmod_zadd1_eq_raw:  | 
| 46820 | 1383  | 
"[|a \<in> int; b \<in> int; c \<in> int|]  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1384  | 
==> (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
 | 
1385  | 
apply (case_tac "c = #0")  | 
| 46820 | 1386  | 
apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD)  | 
1387  | 
apply (blast intro: zadd1_lemma [OF quorem_div_mod quorem_div_mod,  | 
|
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1388  | 
THEN quorem_mod])  | 
| 
 
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 zmod_zadd1_eq: "(a$+b) zmod c = (a zmod c $+ b zmod c) zmod c"  | 
| 46820 | 1392  | 
apply (cut_tac a = "intify (a)" and b = "intify (b)" and c = "intify (c)"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1393  | 
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
 | 
1394  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1395  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1396  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1397  | 
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
 | 
1398  | 
"[|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
 | 
1399  | 
apply (case_tac "b = #0")  | 
| 46820 | 1400  | 
apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD)  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1401  | 
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
 | 
1402  | 
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
 | 
1403  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1404  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1405  | 
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
 | 
1406  | 
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
 | 
1407  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1408  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1409  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1410  | 
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
 | 
1411  | 
"[|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
 | 
1412  | 
apply (case_tac "b = #0")  | 
| 46820 | 1413  | 
apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD)  | 
1414  | 
apply (auto simp add: neq_iff_zless pos_mod_sign pos_mod_bound  | 
|
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1415  | 
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
 | 
1416  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1417  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1418  | 
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
 | 
1419  | 
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
 | 
1420  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1421  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1422  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1423  | 
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
 | 
1424  | 
apply (rule trans [symmetric])  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1425  | 
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
 | 
1426  | 
apply (simp (no_asm))  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1427  | 
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
 | 
1428  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1429  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1430  | 
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
 | 
1431  | 
apply (rule trans [symmetric])  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1432  | 
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
 | 
1433  | 
apply (simp (no_asm))  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1434  | 
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
 | 
1435  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1436  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1437  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1438  | 
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
 | 
1439  | 
"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
 | 
1440  | 
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
 | 
1441  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1442  | 
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
 | 
1443  | 
"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
 | 
1444  | 
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
 | 
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_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
 | 
1447  | 
apply (case_tac "a = #0")  | 
| 46820 | 1448  | 
apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD)  | 
| 
26056
 
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_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
 | 
1450  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1451  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1452  | 
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
 | 
1453  | 
apply (case_tac "a = #0")  | 
| 46820 | 1454  | 
apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD)  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1455  | 
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
 | 
1456  | 
done  | 
| 
 
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  | 
|
| 60770 | 1459  | 
subsection\<open>proving a zdiv (b*c) = (a zdiv b) zdiv c\<close>  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1460  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1461  | 
(*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
 | 
1462  | 
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
 | 
1463  | 
to cause particular problems.*)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1464  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1465  | 
(** 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
 | 
1466  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1467  | 
lemma zdiv_zmult2_aux1:  | 
| 61395 | 1468  | 
"[| #0 $< c; b $< r; r $\<le> #0 |] ==> b$*c $< b$*(q zmod c) $+ r"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1469  | 
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
 | 
1470  | 
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
 | 
1471  | 
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
 | 
1472  | 
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
 | 
1473  | 
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
 | 
1474  | 
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
 | 
1475  | 
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
 | 
1476  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1477  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1478  | 
lemma zdiv_zmult2_aux2:  | 
| 61395 | 1479  | 
"[| #0 $< c; b $< r; r $\<le> #0 |] ==> b $* (q zmod c) $+ r $\<le> #0"  | 
1480  | 
apply (subgoal_tac "b $* (q zmod c) $\<le> #0")  | 
|
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1481  | 
prefer 2  | 
| 46820 | 1482  | 
apply (simp add: zmult_le_0_iff pos_mod_sign)  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1483  | 
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
 | 
1484  | 
(*arithmetic*)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1485  | 
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
 | 
1486  | 
apply assumption  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1487  | 
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
 | 
1488  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1489  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1490  | 
lemma zdiv_zmult2_aux3:  | 
| 61395 | 1491  | 
"[| #0 $< c; #0 $\<le> r; r $< b |] ==> #0 $\<le> b $* (q zmod c) $+ r"  | 
1492  | 
apply (subgoal_tac "#0 $\<le> b $* (q zmod c)")  | 
|
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1493  | 
prefer 2  | 
| 46820 | 1494  | 
apply (simp add: int_0_le_mult_iff pos_mod_sign)  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1495  | 
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
 | 
1496  | 
(*arithmetic*)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1497  | 
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
 | 
1498  | 
apply assumption  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1499  | 
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
 | 
1500  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1501  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1502  | 
lemma zdiv_zmult2_aux4:  | 
| 61395 | 1503  | 
"[| #0 $< c; #0 $\<le> r; r $< b |] ==> b $* (q zmod c) $+ r $< b $* c"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1504  | 
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
 | 
1505  | 
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
 | 
1506  | 
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
 | 
1507  | 
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
 | 
1508  | 
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
 | 
1509  | 
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
 | 
1510  | 
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
 | 
1511  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1512  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1513  | 
lemma zdiv_zmult2_lemma:  | 
| 46820 | 1514  | 
"[| quorem (<a,b>, <q,r>); a \<in> int; b \<in> int; b \<noteq> #0; #0 $< c |]  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1515  | 
==> 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
 | 
1516  | 
apply (auto simp add: zmult_ac zmod_zdiv_equality [symmetric] quorem_def  | 
| 46820 | 1517  | 
neq_iff_zless int_0_less_mult_iff  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1518  | 
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
 | 
1519  | 
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
 | 
1520  | 
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
 | 
1521  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1522  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1523  | 
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
 | 
1524  | 
"[|#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
 | 
1525  | 
apply (case_tac "b = #0")  | 
| 46820 | 1526  | 
apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD)  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1527  | 
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
 | 
1528  | 
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
 | 
1529  | 
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
 | 
1530  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1531  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1532  | 
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
 | 
1533  | 
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
 | 
1534  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1535  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1536  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1537  | 
lemma zmod_zmult2_eq_raw:  | 
| 46820 | 1538  | 
"[|#0 $< c; a \<in> int; b \<in> int|]  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1539  | 
==> 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
 | 
1540  | 
apply (case_tac "b = #0")  | 
| 46820 | 1541  | 
apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD)  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1542  | 
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
 | 
1543  | 
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
 | 
1544  | 
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
 | 
1545  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1546  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1547  | 
lemma zmod_zmult2_eq:  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1548  | 
"#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
 | 
1549  | 
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
 | 
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  | 
|
| 60770 | 1553  | 
subsection\<open>Cancellation of common factors in "zdiv"\<close>  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1554  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1555  | 
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
 | 
1556  | 
"[| #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
 | 
1557  | 
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
 | 
1558  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1559  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1560  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1561  | 
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
 | 
1562  | 
"[| 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
 | 
1563  | 
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
 | 
1564  | 
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
 | 
1565  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1566  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1567  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1568  | 
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
 | 
1569  | 
"[|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
 | 
1570  | 
apply (case_tac "b = #0")  | 
| 46820 | 1571  | 
apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD)  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1572  | 
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
 | 
1573  | 
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
 | 
1574  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1575  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1576  | 
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
 | 
1577  | 
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
 | 
1578  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1579  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1580  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1581  | 
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
 | 
1582  | 
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
 | 
1583  | 
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
 | 
1584  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1585  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1586  | 
|
| 60770 | 1587  | 
subsection\<open>Distribution of factors over "zmod"\<close>  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1588  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1589  | 
lemma zmod_zmult_zmult1_aux1:  | 
| 46820 | 1590  | 
"[| #0 $< b; intify(c) \<noteq> #0 |]  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1591  | 
==> (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
 | 
1592  | 
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
 | 
1593  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1594  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1595  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1596  | 
lemma zmod_zmult_zmult1_aux2:  | 
| 46820 | 1597  | 
"[| b $< #0; intify(c) \<noteq> #0 |]  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1598  | 
==> (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
 | 
1599  | 
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
 | 
1600  | 
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
 | 
1601  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1602  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1603  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1604  | 
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
 | 
1605  | 
"[|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
 | 
1606  | 
apply (case_tac "b = #0")  | 
| 46820 | 1607  | 
apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD)  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1608  | 
apply (case_tac "c = #0")  | 
| 46820 | 1609  | 
apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD)  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1610  | 
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
 | 
1611  | 
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
 | 
1612  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1613  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1614  | 
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
 | 
1615  | 
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
 | 
1616  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1617  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1618  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1619  | 
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
 | 
1620  | 
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
 | 
1621  | 
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
 | 
1622  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1623  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1624  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1625  | 
(** Quotients of signs **)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1626  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1627  | 
lemma zdiv_neg_pos_less0: "[| a $< #0; #0 $< b |] ==> a zdiv b $< #0"  | 
| 61395 | 1628  | 
apply (subgoal_tac "a zdiv b $\<le> #-1")  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1629  | 
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
 | 
1630  | 
apply (simp (no_asm))  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1631  | 
apply (rule zle_trans)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1632  | 
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
 | 
1633  | 
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
 | 
1634  | 
apply (simp (no_asm))  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1635  | 
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
 | 
1636  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1637  | 
|
| 61395 | 1638  | 
lemma zdiv_nonneg_neg_le0: "[| #0 $\<le> a; b $< #0 |] ==> a zdiv b $\<le> #0"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1639  | 
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
 | 
1640  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1641  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1642  | 
|
| 61395 | 1643  | 
lemma pos_imp_zdiv_nonneg_iff: "#0 $< b ==> (#0 $\<le> a zdiv b) \<longleftrightarrow> (#0 $\<le> a)"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1644  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1645  | 
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
 | 
1646  | 
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
 | 
1647  | 
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
 | 
1648  | 
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
 | 
1649  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1650  | 
|
| 61395 | 1651  | 
lemma neg_imp_zdiv_nonneg_iff: "b $< #0 ==> (#0 $\<le> a zdiv b) \<longleftrightarrow> (a $\<le> #0)"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1652  | 
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
 | 
1653  | 
apply (rule iff_trans)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1654  | 
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
 | 
1655  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1656  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1657  | 
|
| 61395 | 1658  | 
(*But not (a zdiv b $\<le> 0 iff a$\<le>0); consider a=1, b=2 when a zdiv b = 0.*)  | 
| 
46821
 
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
 
paulson 
parents: 
46820 
diff
changeset
 | 
1659  | 
lemma pos_imp_zdiv_neg_iff: "#0 $< b ==> (a zdiv b $< #0) \<longleftrightarrow> (a $< #0)"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1660  | 
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
 | 
1661  | 
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
 | 
1662  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1663  | 
|
| 61395 | 1664  | 
(*Again the law fails for $\<le>: consider a = -1, b = -2 when a zdiv b = 0*)  | 
| 
46821
 
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
 
paulson 
parents: 
46820 
diff
changeset
 | 
1665  | 
lemma neg_imp_zdiv_neg_iff: "b $< #0 ==> (a zdiv b $< #0) \<longleftrightarrow> (#0 $< a)"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1666  | 
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
 | 
1667  | 
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
 | 
1668  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1669  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1670  | 
(*  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1671  | 
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
 | 
1672  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1673  | 
 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
 | 
1674  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1675  | 
(** computing "zdiv" by shifting **)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1676  | 
|
| 61395 | 1677  | 
lemma pos_zdiv_mult_2: "#0 $\<le> a ==> (#1 $+ #2$*b) zdiv (#2$*a) = b zdiv a"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1678  | 
apply (case_tac "a = #0")  | 
| 61395 | 1679  | 
apply (subgoal_tac "#1 $\<le> a")  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1680  | 
apply (arith_tac 2)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1681  | 
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
 | 
1682  | 
apply (arith_tac 2)  | 
| 61395 | 1683  | 
apply (subgoal_tac "#2$* (#1 $+ b zmod a) $\<le> #2$*a")  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1684  | 
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
 | 
1685  | 
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
 | 
1686  | 
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
 | 
1687  | 
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
 | 
1688  | 
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
 | 
1689  | 
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
 | 
1690  | 
apply (auto simp add: zmod_pos_pos_trivial)  | 
| 61395 | 1691  | 
apply (subgoal_tac "#0 $\<le> b zmod a")  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1692  | 
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
 | 
1693  | 
apply arith  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1694  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1695  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1696  | 
|
| 61395 | 1697  | 
lemma neg_zdiv_mult_2: "a $\<le> #0 ==> (#1 $+ #2$*b) zdiv (#2$*a) \<longleftrightarrow> (b$+#1) zdiv a"  | 
| 
46821
 
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
 
paulson 
parents: 
46820 
diff
changeset
 | 
1698  | 
apply (subgoal_tac " (#1 $+ #2$* ($-b-#1)) zdiv (#2 $* ($-a)) \<longleftrightarrow> ($-b-#1) zdiv ($-a)")  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1699  | 
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
 | 
1700  | 
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
 | 
1701  | 
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
 | 
1702  | 
apply (Simp_tac 2)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1703  | 
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
 | 
1704  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1705  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1706  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1707  | 
(*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
 | 
1708  | 
simplification problems*)  | 
| 61395 | 1709  | 
lemma lemma: "~ #0 $\<le> x ==> x $\<le> #0"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1710  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1711  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1712  | 
|
| 46820 | 1713  | 
lemma zdiv_integ_of_BIT: "integ_of (v BIT b) zdiv integ_of (w BIT False) =  | 
| 61395 | 1714  | 
(if ~b | #0 $\<le> integ_of w  | 
| 46820 | 1715  | 
then integ_of v zdiv (integ_of w)  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1716  | 
else (integ_of v $+ #1) zdiv (integ_of w))"  | 
| 51686 | 1717  | 
 apply (simp_tac (simpset_of @{theory_context Int} 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
 | 
1718  | 
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
 | 
1719  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1720  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1721  | 
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
 | 
1722  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1723  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1724  | 
(** 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
 | 
1725  | 
|
| 61395 | 1726  | 
lemma pos_zmod_mult_2: "#0 $\<le> a ==> (#1 $+ #2$*b) zmod (#2$*a) = #1 $+ #2 $* (b zmod a)"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1727  | 
apply (case_tac "a = #0")  | 
| 61395 | 1728  | 
apply (subgoal_tac "#1 $\<le> a")  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1729  | 
apply (arith_tac 2)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1730  | 
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
 | 
1731  | 
apply (arith_tac 2)  | 
| 61395 | 1732  | 
apply (subgoal_tac "#2$* (#1 $+ b zmod a) $\<le> #2$*a")  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1733  | 
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
 | 
1734  | 
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
 | 
1735  | 
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
 | 
1736  | 
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
 | 
1737  | 
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
 | 
1738  | 
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
 | 
1739  | 
apply (auto simp add: zmod_pos_pos_trivial)  | 
| 61395 | 1740  | 
apply (subgoal_tac "#0 $\<le> b zmod a")  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1741  | 
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
 | 
1742  | 
apply arith  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1743  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1744  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1745  | 
|
| 61395 | 1746  | 
lemma neg_zmod_mult_2: "a $\<le> #0 ==> (#1 $+ #2$*b) zmod (#2$*a) = #2 $* ((b$+#1) zmod a) - #1"  | 
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1747  | 
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
 | 
1748  | 
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
 | 
1749  | 
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
 | 
1750  | 
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
 | 
1751  | 
apply (Simp_tac 2)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1752  | 
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
 | 
1753  | 
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
 | 
1754  | 
apply auto  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1755  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1756  | 
|
| 46820 | 1757  | 
lemma zmod_integ_of_BIT: "integ_of (v BIT b) zmod integ_of (w BIT False) =  | 
1758  | 
(if b then  | 
|
| 61395 | 1759  | 
if #0 $\<le> integ_of w  | 
| 46820 | 1760  | 
then #2 $* (integ_of v zmod integ_of w) $+ #1  | 
1761  | 
else #2 $* ((integ_of v $+ #1) zmod integ_of w) - #1  | 
|
| 
26056
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1762  | 
else #2 $* (integ_of v zmod integ_of w))"  | 
| 51686 | 1763  | 
 apply (simp_tac (simpset_of @{theory_context Int} 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
 | 
1764  | 
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
 | 
1765  | 
done  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1766  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1767  | 
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
 | 
1768  | 
*)  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1769  | 
|
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1770  | 
end  | 
| 
 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 
krauss 
parents:  
diff
changeset
 | 
1771  |