src/ZF/Integ/IntDiv.thy
author wenzelm
Wed, 14 Nov 2001 23:19:09 +0100
changeset 12189 4729bbf86626
parent 11871 0493188cff42
child 13520 a3d5d8b03d63
permissions -rw-r--r--
Isar attribute and method setup;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
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
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
     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: 11321
diff changeset
    17
  adjust :: "[i,i] => i"
0493188cff42 deleted the redundant first argument of adjust(a,b)
paulson
parents: 11321
diff 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: 11321
diff 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
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
    22
(** the division algorithm **)
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
    23
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
    24
constdefs posDivAlg :: "i => i"
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
    25
(*for the case a>=0, b>0*)
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
    26
(*recdef posDivAlg "inv_image less_than (%(a,b). nat_of(a $- b $+ #1))"*)
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
    27
    "posDivAlg(ab) ==
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
    28
       wfrec(measure(int*int, %<a,b>. nat_of (a $- b $+ #1)),
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
    29
	     ab,
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
    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: 11321
diff changeset
    31
                       else adjust(b, f ` <a,#2$*b>))"
9955
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
    32
11321
01cbbf33779b the rest of integer division
paulson
parents: 9955
diff changeset
    33
01cbbf33779b the rest of integer division
paulson
parents: 9955
diff changeset
    34
(*for the case a<0, b>0*)
01cbbf33779b the rest of integer division
paulson
parents: 9955
diff changeset
    35
constdefs negDivAlg :: "i => i"
01cbbf33779b the rest of integer division
paulson
parents: 9955
diff changeset
    36
(*recdef negDivAlg "inv_image less_than (%(a,b). nat_of(- a $- b))"*)
01cbbf33779b the rest of integer division
paulson
parents: 9955
diff changeset
    37
    "negDivAlg(ab) ==
01cbbf33779b the rest of integer division
paulson
parents: 9955
diff changeset
    38
       wfrec(measure(int*int, %<a,b>. nat_of ($- a $- b)),
01cbbf33779b the rest of integer division
paulson
parents: 9955
diff changeset
    39
	     ab,
01cbbf33779b the rest of integer division
paulson
parents: 9955
diff changeset
    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: 11321
diff changeset
    41
                       else adjust(b, f ` <a,#2$*b>))"
11321
01cbbf33779b the rest of integer division
paulson
parents: 9955
diff changeset
    42
01cbbf33779b the rest of integer division
paulson
parents: 9955
diff changeset
    43
(*for the general case b\\<noteq>0*)
01cbbf33779b the rest of integer division
paulson
parents: 9955
diff changeset
    44
01cbbf33779b the rest of integer division
paulson
parents: 9955
diff changeset
    45
constdefs
01cbbf33779b the rest of integer division
paulson
parents: 9955
diff changeset
    46
  negateSnd :: "i => i"
01cbbf33779b the rest of integer division
paulson
parents: 9955
diff changeset
    47
    "negateSnd == %<q,r>. <q, $-r>"
01cbbf33779b the rest of integer division
paulson
parents: 9955
diff changeset
    48
01cbbf33779b the rest of integer division
paulson
parents: 9955
diff changeset
    49
  (*The full division algorithm considers all possible signs for a, b
01cbbf33779b the rest of integer division
paulson
parents: 9955
diff changeset
    50
    including the special case a=0, b<0, because negDivAlg requires a<0*)
01cbbf33779b the rest of integer division
paulson
parents: 9955
diff changeset
    51
  divAlg :: "i => i"
01cbbf33779b the rest of integer division
paulson
parents: 9955
diff changeset
    52
    "divAlg ==
01cbbf33779b the rest of integer division
paulson
parents: 9955
diff changeset
    53
       %<a,b>. if #0 $<= a then
01cbbf33779b the rest of integer division
paulson
parents: 9955
diff changeset
    54
                  if #0 $<= b then posDivAlg (<a,b>)
01cbbf33779b the rest of integer division
paulson
parents: 9955
diff changeset
    55
                  else if a=#0 then <#0,#0>
01cbbf33779b the rest of integer division
paulson
parents: 9955
diff changeset
    56
                       else negateSnd (negDivAlg (<$-a,$-b>))
01cbbf33779b the rest of integer division
paulson
parents: 9955
diff changeset
    57
               else 
01cbbf33779b the rest of integer division
paulson
parents: 9955
diff changeset
    58
                  if #0$<b then negDivAlg (<a,b>)
01cbbf33779b the rest of integer division
paulson
parents: 9955
diff changeset
    59
                  else         negateSnd (posDivAlg (<$-a,$-b>))"
01cbbf33779b the rest of integer division
paulson
parents: 9955
diff changeset
    60
01cbbf33779b the rest of integer division
paulson
parents: 9955
diff changeset
    61
  zdiv  :: [i,i]=>i                    (infixl "zdiv" 70) 
01cbbf33779b the rest of integer division
paulson
parents: 9955
diff changeset
    62
    "a zdiv b == fst (divAlg (<intify(a), intify(b)>))"
01cbbf33779b the rest of integer division
paulson
parents: 9955
diff changeset
    63
01cbbf33779b the rest of integer division
paulson
parents: 9955
diff changeset
    64
  zmod  :: [i,i]=>i                    (infixl "zmod" 70)
01cbbf33779b the rest of integer division
paulson
parents: 9955
diff changeset
    65
    "a zmod b == snd (divAlg (<intify(a), intify(b)>))"
9955
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
    66
9578
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    67
end