src/HOL/Real/PRat.thy
author wenzelm
Tue, 24 Aug 1999 11:50:58 +0200
changeset 7333 6cb15c6f1d9f
parent 7219 4e3f386c2e37
child 7376 46f92a120af9
permissions -rw-r--r--
isar: no_pos flag;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
     1
(*  Title       : PRat.thy
7219
4e3f386c2e37 inserted Id: lines
paulson
parents: 7077
diff changeset
     2
    ID          : $Id$
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
     3
    Author      : Jacques D. Fleuriot
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
     4
    Copyright   : 1998  University of Cambridge
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
     5
    Description : The positive rationals
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
     6
*) 
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
     7
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
     8
PRat = PNat + Equiv +
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
     9
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    10
constdefs
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    11
    ratrel   ::  "((pnat * pnat) * (pnat * pnat)) set"
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    12
    "ratrel  ==  {p. ? x1 y1 x2 y2. p=((x1::pnat,y1),(x2,y2)) & x1*y2 = x2*y1}" 
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    13
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    14
typedef prat = "{x::(pnat*pnat).True}/ratrel"          (Equiv.quotient_def)
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    15
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    16
instance
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    17
   prat  :: {ord,plus,times}
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    18
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    19
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    20
constdefs
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    21
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5078
diff changeset
    22
  prat_of_pnat :: pnat => prat           
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5078
diff changeset
    23
  "prat_of_pnat m     == Abs_prat(ratrel^^{(m,Abs_pnat 1)})"
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    24
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    25
  qinv      :: prat => prat
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5078
diff changeset
    26
  "qinv(Q)  == Abs_prat(UN (x,y):Rep_prat(Q). ratrel^^{(y,x)})" 
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    27
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    28
defs
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    29
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    30
  prat_add_def  
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    31
  "P + Q == Abs_prat(UN p1:Rep_prat(P). UN p2:Rep_prat(Q).
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    32
                split(%x1 y1. split(%x2 y2. ratrel^^{(x1*y2 + x2*y1, y1*y2)}) p2) p1)"
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    33
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    34
  prat_mult_def  
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    35
  "P * Q == Abs_prat(UN p1:Rep_prat(P). UN p2:Rep_prat(Q).
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    36
                split(%x1 y1. split(%x2 y2. ratrel^^{(x1*x2, y1*y2)}) p2) p1)"
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    37
 
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    38
  (*** Gleason p. 119 ***)
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    39
  prat_less_def
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    40
  "P < (Q::prat) == ? T. P + T = Q"
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    41
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    42
  prat_le_def
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    43
  "P <= (Q::prat) == ~(Q < P)" 
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    44
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    45
end
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    46
  
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    47
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    48
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    49
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    50