| author | paulson | 
| Fri, 24 May 2002 16:55:28 +0200 | |
| changeset 13178 | bc54319f6875 | 
| parent 11871 | 0493188cff42 | 
| child 13520 | a3d5d8b03d63 | 
| permissions | -rw-r--r-- | 
| 9578 
ab26d6c8ebfe
new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
 paulson parents: diff
changeset | 1 | (* Title: ZF/IntDiv.thy | 
| 
ab26d6c8ebfe
new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
 paulson parents: diff
changeset | 2 | ID: $Id$ | 
| 
ab26d6c8ebfe
new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
 paulson parents: diff
changeset | 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 
ab26d6c8ebfe
new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
 paulson parents: diff
changeset | 4 | Copyright 1999 University of Cambridge | 
| 
ab26d6c8ebfe
new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
 paulson parents: diff
changeset | 5 | |
| 
ab26d6c8ebfe
new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
 paulson parents: diff
changeset | 6 | The division operators div, mod | 
| 
ab26d6c8ebfe
new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
 paulson parents: diff
changeset | 7 | *) | 
| 
ab26d6c8ebfe
new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
 paulson parents: diff
changeset | 8 | |
| 9955 | 9 | IntDiv = IntArith + OrderArith + | 
| 9578 
ab26d6c8ebfe
new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
 paulson parents: diff
changeset | 10 | |
| 
ab26d6c8ebfe
new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
 paulson parents: diff
changeset | 11 | constdefs | 
| 
ab26d6c8ebfe
new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
 paulson parents: diff
changeset | 12 | quorem :: "[i,i] => o" | 
| 
ab26d6c8ebfe
new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
 paulson parents: diff
changeset | 13 | "quorem == %<a,b> <q,r>. | 
| 
ab26d6c8ebfe
new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
 paulson parents: diff
changeset | 14 | a = b$*q $+ r & | 
| 
ab26d6c8ebfe
new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
 paulson parents: diff
changeset | 15 | (#0$<b & #0$<=r & r$<b | ~(#0$<b) & b$<r & r $<= #0)" | 
| 
ab26d6c8ebfe
new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
 paulson parents: diff
changeset | 16 | |
| 11871 
0493188cff42
deleted the redundant first argument of adjust(a,b)
 paulson parents: 
11321diff
changeset | 17 | adjust :: "[i,i] => i" | 
| 
0493188cff42
deleted the redundant first argument of adjust(a,b)
 paulson parents: 
11321diff
changeset | 18 | "adjust(b) == %<q,r>. if #0 $<= r$-b then <#2$*q $+ #1,r$-b> | 
| 
0493188cff42
deleted the redundant first argument of adjust(a,b)
 paulson parents: 
11321diff
changeset | 19 | else <#2$*q,r>" | 
| 9578 
ab26d6c8ebfe
new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
 paulson parents: diff
changeset | 20 | |
| 
ab26d6c8ebfe
new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
 paulson parents: diff
changeset | 21 | |
| 9955 | 22 | (** the division algorithm **) | 
| 23 | ||
| 24 | constdefs posDivAlg :: "i => i" | |
| 25 | (*for the case a>=0, b>0*) | |
| 26 | (*recdef posDivAlg "inv_image less_than (%(a,b). nat_of(a $- b $+ #1))"*) | |
| 27 | "posDivAlg(ab) == | |
| 28 | wfrec(measure(int*int, %<a,b>. nat_of (a $- b $+ #1)), | |
| 29 | ab, | |
| 30 | %<a,b> f. if (a$<b | b$<=#0) then <#0,a> | |
| 11871 
0493188cff42
deleted the redundant first argument of adjust(a,b)
 paulson parents: 
11321diff
changeset | 31 | else adjust(b, f ` <a,#2$*b>))" | 
| 9955 | 32 | |
| 11321 | 33 | |
| 34 | (*for the case a<0, b>0*) | |
| 35 | constdefs negDivAlg :: "i => i" | |
| 36 | (*recdef negDivAlg "inv_image less_than (%(a,b). nat_of(- a $- b))"*) | |
| 37 | "negDivAlg(ab) == | |
| 38 | wfrec(measure(int*int, %<a,b>. nat_of ($- a $- b)), | |
| 39 | ab, | |
| 40 | %<a,b> f. if (#0 $<= a$+b | b$<=#0) then <#-1,a$+b> | |
| 11871 
0493188cff42
deleted the redundant first argument of adjust(a,b)
 paulson parents: 
11321diff
changeset | 41 | else adjust(b, f ` <a,#2$*b>))" | 
| 11321 | 42 | |
| 43 | (*for the general case b\\<noteq>0*) | |
| 44 | ||
| 45 | constdefs | |
| 46 | negateSnd :: "i => i" | |
| 47 | "negateSnd == %<q,r>. <q, $-r>" | |
| 48 | ||
| 49 | (*The full division algorithm considers all possible signs for a, b | |
| 50 | including the special case a=0, b<0, because negDivAlg requires a<0*) | |
| 51 | divAlg :: "i => i" | |
| 52 | "divAlg == | |
| 53 | %<a,b>. if #0 $<= a then | |
| 54 | if #0 $<= b then posDivAlg (<a,b>) | |
| 55 | else if a=#0 then <#0,#0> | |
| 56 | else negateSnd (negDivAlg (<$-a,$-b>)) | |
| 57 | else | |
| 58 | if #0$<b then negDivAlg (<a,b>) | |
| 59 | else negateSnd (posDivAlg (<$-a,$-b>))" | |
| 60 | ||
| 61 | zdiv :: [i,i]=>i (infixl "zdiv" 70) | |
| 62 | "a zdiv b == fst (divAlg (<intify(a), intify(b)>))" | |
| 63 | ||
| 64 | zmod :: [i,i]=>i (infixl "zmod" 70) | |
| 65 | "a zmod b == snd (divAlg (<intify(a), intify(b)>))" | |
| 9955 | 66 | |
| 9578 
ab26d6c8ebfe
new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
 paulson parents: diff
changeset | 67 | end |