clasohm@0: (* Title: ZF/ex/integ.thy clasohm@0: ID: $Id$ clasohm@0: Author: Lawrence C Paulson, Cambridge University Computer Laboratory clasohm@0: Copyright 1993 University of Cambridge clasohm@0: clasohm@0: The integers as equivalence classes over nat*nat. clasohm@0: *) clasohm@0: lcp@532: Integ = EquivClass + Arith + clasohm@0: consts clasohm@0: intrel,integ:: "i" clasohm@0: znat :: "i=>i" ("$# _" [80] 80) clasohm@0: zminus :: "i=>i" ("$~ _" [80] 80) clasohm@0: znegative :: "i=>o" clasohm@0: zmagnitude :: "i=>i" clasohm@0: "$*" :: "[i,i]=>i" (infixl 70) clasohm@0: "$'/" :: "[i,i]=>i" (infixl 70) clasohm@0: "$'/'/" :: "[i,i]=>i" (infixl 70) clasohm@0: "$+" :: "[i,i]=>i" (infixl 65) clasohm@0: "$-" :: "[i,i]=>i" (infixl 65) clasohm@0: "$<" :: "[i,i]=>o" (infixl 50) clasohm@0: clasohm@0: rules clasohm@0: clasohm@0: intrel_def clasohm@0: "intrel == {p:(nat*nat)*(nat*nat). \ clasohm@0: \ EX x1 y1 x2 y2. p=<,> & x1#+y2 = x2#+y1}" clasohm@0: clasohm@0: integ_def "integ == (nat*nat)/intrel" clasohm@0: clasohm@0: znat_def "$# m == intrel `` {}" clasohm@0: clasohm@0: zminus_def "$~ Z == UN p:Z. split(%x y. intrel``{}, p)" clasohm@0: clasohm@0: znegative_def lcp@29: "znegative(Z) == EX x y. x:Z" clasohm@0: clasohm@0: zmagnitude_def clasohm@0: "zmagnitude(Z) == UN p:Z. split(%x y. (y#-x) #+ (x#-y), p)" clasohm@0: clasohm@0: zadd_def clasohm@0: "Z1 $+ Z2 == \ clasohm@0: \ UN p1:Z1. UN p2:Z2. split(%x1 y1. split(%x2 y2. \ clasohm@0: \ intrel``{}, p2), p1)" clasohm@0: clasohm@0: zdiff_def "Z1 $- Z2 == Z1 $+ zminus(Z2)" clasohm@0: zless_def "Z1 $< Z2 == znegative(Z1 $- Z2)" clasohm@0: clasohm@0: zmult_def clasohm@0: "Z1 $* Z2 == \ clasohm@0: \ UN p1:Z1. UN p2:Z2. split(%x1 y1. split(%x2 y2. \ clasohm@0: \ intrel``{}, p2), p1)" clasohm@0: clasohm@0: end