src/HOL/Real/Rational.thy
author paulson
Sun, 15 Feb 2004 10:46:37 +0100
changeset 14387 e96d5c42c4b0
parent 14378 69c4d5997669
child 14421 ee97b6463cb4
permissions -rw-r--r--
Polymorphic treatment of binary arithmetic using axclasses
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
     1
(*  Title: HOL/Library/Rational.thy
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
     2
    ID:    $Id$
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
     3
    Author: Markus Wenzel, TU Muenchen
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
     4
    License: GPL (GNU GENERAL PUBLIC LICENSE)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
     5
*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
     6
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
     7
header {*
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
     8
  \title{Rational numbers}
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
     9
  \author{Markus Wenzel}
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    10
*}
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    11
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    12
theory Rational = Quotient + Main
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    13
files ("rat_arith.ML"):
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    14
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    15
subsection {* Fractions *}
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    16
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    17
subsubsection {* The type of fractions *}
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    18
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    19
typedef fraction = "{(a, b) :: int \<times> int | a b. b \<noteq> 0}"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    20
proof
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    21
  show "(0, 1) \<in> ?fraction" by simp
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    22
qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    23
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    24
constdefs
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    25
  fract :: "int => int => fraction"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    26
  "fract a b == Abs_fraction (a, b)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    27
  num :: "fraction => int"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    28
  "num Q == fst (Rep_fraction Q)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    29
  den :: "fraction => int"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    30
  "den Q == snd (Rep_fraction Q)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    31
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    32
lemma fract_num [simp]: "b \<noteq> 0 ==> num (fract a b) = a"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    33
  by (simp add: fract_def num_def fraction_def Abs_fraction_inverse)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    34
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    35
lemma fract_den [simp]: "b \<noteq> 0 ==> den (fract a b) = b"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    36
  by (simp add: fract_def den_def fraction_def Abs_fraction_inverse)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    37
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    38
lemma fraction_cases [case_names fract, cases type: fraction]:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    39
  "(!!a b. Q = fract a b ==> b \<noteq> 0 ==> C) ==> C"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    40
proof -
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    41
  assume r: "!!a b. Q = fract a b ==> b \<noteq> 0 ==> C"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    42
  obtain a b where "Q = fract a b" and "b \<noteq> 0"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    43
    by (cases Q) (auto simp add: fract_def fraction_def)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    44
  thus C by (rule r)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    45
qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    46
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    47
lemma fraction_induct [case_names fract, induct type: fraction]:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    48
    "(!!a b. b \<noteq> 0 ==> P (fract a b)) ==> P Q"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    49
  by (cases Q) simp
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    50
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    51
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    52
subsubsection {* Equivalence of fractions *}
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    53
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    54
instance fraction :: eqv ..
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    55
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    56
defs (overloaded)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    57
  equiv_fraction_def: "Q \<sim> R == num Q * den R = num R * den Q"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    58
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    59
lemma equiv_fraction_iff [iff]:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    60
    "b \<noteq> 0 ==> b' \<noteq> 0 ==> (fract a b \<sim> fract a' b') = (a * b' = a' * b)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    61
  by (simp add: equiv_fraction_def)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    62
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    63
instance fraction :: equiv
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    64
proof
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    65
  fix Q R S :: fraction
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    66
  {
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    67
    show "Q \<sim> Q"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    68
    proof (induct Q)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    69
      fix a b :: int
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    70
      assume "b \<noteq> 0" and "b \<noteq> 0"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    71
      with refl show "fract a b \<sim> fract a b" ..
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    72
    qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    73
  next
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    74
    assume "Q \<sim> R" and "R \<sim> S"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    75
    show "Q \<sim> S"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    76
    proof (insert prems, induct Q, induct R, induct S)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    77
      fix a b a' b' a'' b'' :: int
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    78
      assume b: "b \<noteq> 0" and b': "b' \<noteq> 0" and b'': "b'' \<noteq> 0"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    79
      assume "fract a b \<sim> fract a' b'" hence eq1: "a * b' = a' * b" ..
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    80
      assume "fract a' b' \<sim> fract a'' b''" hence eq2: "a' * b'' = a'' * b'" ..
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    81
      have "a * b'' = a'' * b"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    82
      proof cases
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    83
        assume "a' = 0"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    84
        with b' eq1 eq2 have "a = 0 \<and> a'' = 0" by auto
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    85
        thus ?thesis by simp
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    86
      next
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    87
        assume a': "a' \<noteq> 0"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    88
        from eq1 eq2 have "(a * b') * (a' * b'') = (a' * b) * (a'' * b')" by simp
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    89
        hence "(a * b'') * (a' * b') = (a'' * b) * (a' * b')" by (simp only: mult_ac)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    90
        with a' b' show ?thesis by simp
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    91
      qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    92
      thus "fract a b \<sim> fract a'' b''" ..
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    93
    qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    94
  next
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    95
    show "Q \<sim> R ==> R \<sim> Q"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    96
    proof (induct Q, induct R)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    97
      fix a b a' b' :: int
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    98
      assume b: "b \<noteq> 0" and b': "b' \<noteq> 0"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    99
      assume "fract a b \<sim> fract a' b'"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   100
      hence "a * b' = a' * b" ..
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   101
      hence "a' * b = a * b'" ..
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   102
      thus "fract a' b' \<sim> fract a b" ..
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   103
    qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   104
  }
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   105
qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   106
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   107
lemma eq_fraction_iff [iff]:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   108
    "b \<noteq> 0 ==> b' \<noteq> 0 ==> (\<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor>) = (a * b' = a' * b)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   109
  by (simp add: equiv_fraction_iff quot_equality)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   110
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   111
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   112
subsubsection {* Operations on fractions *}
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   113
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   114
text {*
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   115
 We define the basic arithmetic operations on fractions and
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   116
 demonstrate their ``well-definedness'', i.e.\ congruence with respect
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   117
 to equivalence of fractions.
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   118
*}
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   119
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   120
instance fraction :: zero ..
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   121
instance fraction :: one ..
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   122
instance fraction :: plus ..
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   123
instance fraction :: minus ..
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   124
instance fraction :: times ..
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   125
instance fraction :: inverse ..
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   126
instance fraction :: ord ..
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   127
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   128
defs (overloaded)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   129
  zero_fraction_def: "0 == fract 0 1"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   130
  one_fraction_def: "1 == fract 1 1"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   131
  add_fraction_def: "Q + R ==
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   132
    fract (num Q * den R + num R * den Q) (den Q * den R)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   133
  minus_fraction_def: "-Q == fract (-(num Q)) (den Q)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   134
  mult_fraction_def: "Q * R == fract (num Q * num R) (den Q * den R)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   135
  inverse_fraction_def: "inverse Q == fract (den Q) (num Q)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   136
  le_fraction_def: "Q \<le> R ==
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   137
    (num Q * den R) * (den Q * den R) \<le> (num R * den Q) * (den Q * den R)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   138
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   139
lemma is_zero_fraction_iff: "b \<noteq> 0 ==> (\<lfloor>fract a b\<rfloor> = \<lfloor>0\<rfloor>) = (a = 0)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   140
  by (simp add: zero_fraction_def eq_fraction_iff)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   141
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   142
theorem add_fraction_cong:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   143
  "\<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor> ==> \<lfloor>fract c d\<rfloor> = \<lfloor>fract c' d'\<rfloor>
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   144
    ==> b \<noteq> 0 ==> b' \<noteq> 0 ==> d \<noteq> 0 ==> d' \<noteq> 0
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   145
    ==> \<lfloor>fract a b + fract c d\<rfloor> = \<lfloor>fract a' b' + fract c' d'\<rfloor>"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   146
proof -
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   147
  assume neq: "b \<noteq> 0"  "b' \<noteq> 0"  "d \<noteq> 0"  "d' \<noteq> 0"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   148
  assume "\<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor>" hence eq1: "a * b' = a' * b" ..
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   149
  assume "\<lfloor>fract c d\<rfloor> = \<lfloor>fract c' d'\<rfloor>" hence eq2: "c * d' = c' * d" ..
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   150
  have "\<lfloor>fract (a * d + c * b) (b * d)\<rfloor> = \<lfloor>fract (a' * d' + c' * b') (b' * d')\<rfloor>"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   151
  proof
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   152
    show "(a * d + c * b) * (b' * d') = (a' * d' + c' * b') * (b * d)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   153
      (is "?lhs = ?rhs")
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   154
    proof -
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   155
      have "?lhs = (a * b') * (d * d') + (c * d') * (b * b')"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   156
        by (simp add: int_distrib mult_ac)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   157
      also have "... = (a' * b) * (d * d') + (c' * d) * (b * b')"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   158
        by (simp only: eq1 eq2)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   159
      also have "... = ?rhs"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   160
        by (simp add: int_distrib mult_ac)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   161
      finally show "?lhs = ?rhs" .
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   162
    qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   163
    from neq show "b * d \<noteq> 0" by simp
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   164
    from neq show "b' * d' \<noteq> 0" by simp
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   165
  qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   166
  with neq show ?thesis by (simp add: add_fraction_def)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   167
qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   168
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   169
theorem minus_fraction_cong:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   170
  "\<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor> ==> b \<noteq> 0 ==> b' \<noteq> 0
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   171
    ==> \<lfloor>-(fract a b)\<rfloor> = \<lfloor>-(fract a' b')\<rfloor>"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   172
proof -
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   173
  assume neq: "b \<noteq> 0"  "b' \<noteq> 0"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   174
  assume "\<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor>"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   175
  hence "a * b' = a' * b" ..
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   176
  hence "-a * b' = -a' * b" by simp
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   177
  hence "\<lfloor>fract (-a) b\<rfloor> = \<lfloor>fract (-a') b'\<rfloor>" ..
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   178
  with neq show ?thesis by (simp add: minus_fraction_def)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   179
qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   180
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   181
theorem mult_fraction_cong:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   182
  "\<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor> ==> \<lfloor>fract c d\<rfloor> = \<lfloor>fract c' d'\<rfloor>
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   183
    ==> b \<noteq> 0 ==> b' \<noteq> 0 ==> d \<noteq> 0 ==> d' \<noteq> 0
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   184
    ==> \<lfloor>fract a b * fract c d\<rfloor> = \<lfloor>fract a' b' * fract c' d'\<rfloor>"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   185
proof -
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   186
  assume neq: "b \<noteq> 0"  "b' \<noteq> 0"  "d \<noteq> 0"  "d' \<noteq> 0"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   187
  assume "\<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor>" hence eq1: "a * b' = a' * b" ..
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   188
  assume "\<lfloor>fract c d\<rfloor> = \<lfloor>fract c' d'\<rfloor>" hence eq2: "c * d' = c' * d" ..
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   189
  have "\<lfloor>fract (a * c) (b * d)\<rfloor> = \<lfloor>fract (a' * c') (b' * d')\<rfloor>"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   190
  proof
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   191
    from eq1 eq2 have "(a * b') * (c * d') = (a' * b) * (c' * d)" by simp
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   192
    thus "(a * c) * (b' * d') = (a' * c') * (b * d)" by (simp add: mult_ac)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   193
    from neq show "b * d \<noteq> 0" by simp
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   194
    from neq show "b' * d' \<noteq> 0" by simp
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   195
  qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   196
  with neq show "\<lfloor>fract a b * fract c d\<rfloor> = \<lfloor>fract a' b' * fract c' d'\<rfloor>"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   197
    by (simp add: mult_fraction_def)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   198
qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   199
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   200
theorem inverse_fraction_cong:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   201
  "\<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor> ==> \<lfloor>fract a b\<rfloor> \<noteq> \<lfloor>0\<rfloor> ==> \<lfloor>fract a' b'\<rfloor> \<noteq> \<lfloor>0\<rfloor>
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   202
    ==> b \<noteq> 0 ==> b' \<noteq> 0
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   203
    ==> \<lfloor>inverse (fract a b)\<rfloor> = \<lfloor>inverse (fract a' b')\<rfloor>"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   204
proof -
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   205
  assume neq: "b \<noteq> 0"  "b' \<noteq> 0"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   206
  assume "\<lfloor>fract a b\<rfloor> \<noteq> \<lfloor>0\<rfloor>" and "\<lfloor>fract a' b'\<rfloor> \<noteq> \<lfloor>0\<rfloor>"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   207
  with neq obtain "a \<noteq> 0" and "a' \<noteq> 0" by (simp add: is_zero_fraction_iff)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   208
  assume "\<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor>"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   209
  hence "a * b' = a' * b" ..
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   210
  hence "b * a' = b' * a" by (simp only: mult_ac)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   211
  hence "\<lfloor>fract b a\<rfloor> = \<lfloor>fract b' a'\<rfloor>" ..
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   212
  with neq show ?thesis by (simp add: inverse_fraction_def)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   213
qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   214
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   215
theorem le_fraction_cong:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   216
  "\<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor> ==> \<lfloor>fract c d\<rfloor> = \<lfloor>fract c' d'\<rfloor>
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   217
    ==> b \<noteq> 0 ==> b' \<noteq> 0 ==> d \<noteq> 0 ==> d' \<noteq> 0
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   218
    ==> (fract a b \<le> fract c d) = (fract a' b' \<le> fract c' d')"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   219
proof -
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   220
  assume neq: "b \<noteq> 0"  "b' \<noteq> 0"  "d \<noteq> 0"  "d' \<noteq> 0"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   221
  assume "\<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor>" hence eq1: "a * b' = a' * b" ..
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   222
  assume "\<lfloor>fract c d\<rfloor> = \<lfloor>fract c' d'\<rfloor>" hence eq2: "c * d' = c' * d" ..
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   223
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   224
  let ?le = "\<lambda>a b c d. ((a * d) * (b * d) \<le> (c * b) * (b * d))"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   225
  {
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   226
    fix a b c d x :: int assume x: "x \<noteq> 0"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   227
    have "?le a b c d = ?le (a * x) (b * x) c d"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   228
    proof -
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   229
      from x have "0 < x * x" by (auto simp add: zero_less_mult_iff)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   230
      hence "?le a b c d =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   231
          ((a * d) * (b * d) * (x * x) \<le> (c * b) * (b * d) * (x * x))"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   232
        by (simp add: mult_le_cancel_right)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   233
      also have "... = ?le (a * x) (b * x) c d"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   234
        by (simp add: mult_ac)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   235
      finally show ?thesis .
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   236
    qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   237
  } note le_factor = this
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   238
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   239
  let ?D = "b * d" and ?D' = "b' * d'"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   240
  from neq have D: "?D \<noteq> 0" by simp
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   241
  from neq have "?D' \<noteq> 0" by simp
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   242
  hence "?le a b c d = ?le (a * ?D') (b * ?D') c d"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   243
    by (rule le_factor)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   244
  also have "... = ((a * b') * ?D * ?D' * d * d' \<le> (c * d') * ?D * ?D' * b * b')"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   245
    by (simp add: mult_ac)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   246
  also have "... = ((a' * b) * ?D * ?D' * d * d' \<le> (c' * d) * ?D * ?D' * b * b')"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   247
    by (simp only: eq1 eq2)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   248
  also have "... = ?le (a' * ?D) (b' * ?D) c' d'"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   249
    by (simp add: mult_ac)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   250
  also from D have "... = ?le a' b' c' d'"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   251
    by (rule le_factor [symmetric])
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   252
  finally have "?le a b c d = ?le a' b' c' d'" .
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   253
  with neq show ?thesis by (simp add: le_fraction_def)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   254
qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   255
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   256
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   257
subsection {* Rational numbers *}
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   258
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   259
subsubsection {* The type of rational numbers *}
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   260
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   261
typedef (Rat)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   262
  rat = "UNIV :: fraction quot set" ..
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   263
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   264
lemma RatI [intro, simp]: "Q \<in> Rat"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   265
  by (simp add: Rat_def)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   266
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   267
constdefs
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   268
  fraction_of :: "rat => fraction"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   269
  "fraction_of q == pick (Rep_Rat q)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   270
  rat_of :: "fraction => rat"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   271
  "rat_of Q == Abs_Rat \<lfloor>Q\<rfloor>"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   272
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   273
theorem rat_of_equality [iff?]: "(rat_of Q = rat_of Q') = (\<lfloor>Q\<rfloor> = \<lfloor>Q'\<rfloor>)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   274
  by (simp add: rat_of_def Abs_Rat_inject)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   275
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   276
lemma rat_of: "\<lfloor>Q\<rfloor> = \<lfloor>Q'\<rfloor> ==> rat_of Q = rat_of Q'" ..
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   277
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   278
constdefs
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   279
  Fract :: "int => int => rat"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   280
  "Fract a b == rat_of (fract a b)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   281
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   282
theorem Fract_inverse: "\<lfloor>fraction_of (Fract a b)\<rfloor> = \<lfloor>fract a b\<rfloor>"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   283
  by (simp add: fraction_of_def rat_of_def Fract_def Abs_Rat_inverse pick_inverse)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   284
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   285
theorem Fract_equality [iff?]:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   286
    "(Fract a b = Fract c d) = (\<lfloor>fract a b\<rfloor> = \<lfloor>fract c d\<rfloor>)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   287
  by (simp add: Fract_def rat_of_equality)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   288
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   289
theorem eq_rat:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   290
    "b \<noteq> 0 ==> d \<noteq> 0 ==> (Fract a b = Fract c d) = (a * d = c * b)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   291
  by (simp add: Fract_equality eq_fraction_iff)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   292
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   293
theorem Rat_cases [case_names Fract, cases type: rat]:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   294
  "(!!a b. q = Fract a b ==> b \<noteq> 0 ==> C) ==> C"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   295
proof -
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   296
  assume r: "!!a b. q = Fract a b ==> b \<noteq> 0 ==> C"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   297
  obtain x where "q = Abs_Rat x" by (cases q)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   298
  moreover obtain Q where "x = \<lfloor>Q\<rfloor>" by (cases x)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   299
  moreover obtain a b where "Q = fract a b" and "b \<noteq> 0" by (cases Q)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   300
  ultimately have "q = Fract a b" by (simp only: Fract_def rat_of_def)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   301
  thus ?thesis by (rule r)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   302
qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   303
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   304
theorem Rat_induct [case_names Fract, induct type: rat]:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   305
    "(!!a b. b \<noteq> 0 ==> P (Fract a b)) ==> P q"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   306
  by (cases q) simp
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   307
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   308
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   309
subsubsection {* Canonical function definitions *}
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   310
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   311
text {*
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   312
  Note that the unconditional version below is much easier to read.
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   313
*}
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   314
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   315
theorem rat_cond_function:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   316
  "(!!q r. P \<lfloor>fraction_of q\<rfloor> \<lfloor>fraction_of r\<rfloor> ==>
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   317
      f q r == g (fraction_of q) (fraction_of r)) ==>
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   318
    (!!a b a' b' c d c' d'.
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   319
      \<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor> ==> \<lfloor>fract c d\<rfloor> = \<lfloor>fract c' d'\<rfloor> ==>
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   320
      P \<lfloor>fract a b\<rfloor> \<lfloor>fract c d\<rfloor> ==> P \<lfloor>fract a' b'\<rfloor> \<lfloor>fract c' d'\<rfloor> ==>
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   321
      b \<noteq> 0 ==> b' \<noteq> 0 ==> d \<noteq> 0 ==> d' \<noteq> 0 ==>
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   322
      g (fract a b) (fract c d) = g (fract a' b') (fract c' d')) ==>
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   323
    P \<lfloor>fract a b\<rfloor> \<lfloor>fract c d\<rfloor> ==>
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   324
      f (Fract a b) (Fract c d) = g (fract a b) (fract c d)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   325
  (is "PROP ?eq ==> PROP ?cong ==> ?P ==> _")
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   326
proof -
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   327
  assume eq: "PROP ?eq" and cong: "PROP ?cong" and P: ?P
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   328
  have "f (Abs_Rat \<lfloor>fract a b\<rfloor>) (Abs_Rat \<lfloor>fract c d\<rfloor>) = g (fract a b) (fract c d)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   329
  proof (rule quot_cond_function)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   330
    fix X Y assume "P X Y"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   331
    with eq show "f (Abs_Rat X) (Abs_Rat Y) == g (pick X) (pick Y)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   332
      by (simp add: fraction_of_def pick_inverse Abs_Rat_inverse)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   333
  next
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   334
    fix Q Q' R R' :: fraction
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   335
    show "\<lfloor>Q\<rfloor> = \<lfloor>Q'\<rfloor> ==> \<lfloor>R\<rfloor> = \<lfloor>R'\<rfloor> ==>
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   336
        P \<lfloor>Q\<rfloor> \<lfloor>R\<rfloor> ==> P \<lfloor>Q'\<rfloor> \<lfloor>R'\<rfloor> ==> g Q R = g Q' R'"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   337
      by (induct Q, induct Q', induct R, induct R') (rule cong)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   338
  qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   339
  thus ?thesis by (unfold Fract_def rat_of_def)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   340
qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   341
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   342
theorem rat_function:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   343
  "(!!q r. f q r == g (fraction_of q) (fraction_of r)) ==>
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   344
    (!!a b a' b' c d c' d'.
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   345
      \<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor> ==> \<lfloor>fract c d\<rfloor> = \<lfloor>fract c' d'\<rfloor> ==>
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   346
      b \<noteq> 0 ==> b' \<noteq> 0 ==> d \<noteq> 0 ==> d' \<noteq> 0 ==>
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   347
      g (fract a b) (fract c d) = g (fract a' b') (fract c' d')) ==>
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   348
    f (Fract a b) (Fract c d) = g (fract a b) (fract c d)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   349
proof -
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   350
  case rule_context from this TrueI
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   351
  show ?thesis by (rule rat_cond_function)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   352
qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   353
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   354
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   355
subsubsection {* Standard operations on rational numbers *}
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   356
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   357
instance rat :: zero ..
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   358
instance rat :: one ..
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   359
instance rat :: plus ..
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   360
instance rat :: minus ..
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   361
instance rat :: times ..
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   362
instance rat :: inverse ..
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   363
instance rat :: ord ..
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   364
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   365
defs (overloaded)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   366
  zero_rat_def: "0 == rat_of 0"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   367
  one_rat_def: "1 == rat_of 1"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   368
  add_rat_def: "q + r == rat_of (fraction_of q + fraction_of r)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   369
  minus_rat_def: "-q == rat_of (-(fraction_of q))"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   370
  diff_rat_def: "q - r == q + (-(r::rat))"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   371
  mult_rat_def: "q * r == rat_of (fraction_of q * fraction_of r)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   372
  inverse_rat_def: "inverse q == 
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   373
                    if q=0 then 0 else rat_of (inverse (fraction_of q))"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   374
  divide_rat_def: "q / r == q * inverse (r::rat)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   375
  le_rat_def: "q \<le> r == fraction_of q \<le> fraction_of r"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   376
  less_rat_def: "q < r == q \<le> r \<and> q \<noteq> (r::rat)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   377
  abs_rat_def: "\<bar>q\<bar> == if q < 0 then -q else (q::rat)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   378
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   379
theorem zero_rat: "0 = Fract 0 1"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   380
  by (simp add: zero_rat_def zero_fraction_def rat_of_def Fract_def)        
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   381
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   382
theorem one_rat: "1 = Fract 1 1"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   383
  by (simp add: one_rat_def one_fraction_def rat_of_def Fract_def)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   384
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   385
theorem add_rat: "b \<noteq> 0 ==> d \<noteq> 0 ==>
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   386
  Fract a b + Fract c d = Fract (a * d + c * b) (b * d)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   387
proof -
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   388
  have "Fract a b + Fract c d = rat_of (fract a b + fract c d)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   389
    by (rule rat_function, rule add_rat_def, rule rat_of, rule add_fraction_cong)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   390
  also
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   391
  assume "b \<noteq> 0"  "d \<noteq> 0"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   392
  hence "fract a b + fract c d = fract (a * d + c * b) (b * d)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   393
    by (simp add: add_fraction_def)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   394
  finally show ?thesis by (unfold Fract_def)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   395
qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   396
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   397
theorem minus_rat: "b \<noteq> 0 ==> -(Fract a b) = Fract (-a) b"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   398
proof -
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   399
  have "-(Fract a b) = rat_of (-(fract a b))"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   400
    by (rule rat_function, rule minus_rat_def, rule rat_of, rule minus_fraction_cong)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   401
  also assume "b \<noteq> 0" hence "-(fract a b) = fract (-a) b"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   402
    by (simp add: minus_fraction_def)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   403
  finally show ?thesis by (unfold Fract_def)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   404
qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   405
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   406
theorem diff_rat: "b \<noteq> 0 ==> d \<noteq> 0 ==>
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   407
    Fract a b - Fract c d = Fract (a * d - c * b) (b * d)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   408
  by (simp add: diff_rat_def add_rat minus_rat)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   409
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   410
theorem mult_rat: "b \<noteq> 0 ==> d \<noteq> 0 ==>
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   411
  Fract a b * Fract c d = Fract (a * c) (b * d)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   412
proof -
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   413
  have "Fract a b * Fract c d = rat_of (fract a b * fract c d)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   414
    by (rule rat_function, rule mult_rat_def, rule rat_of, rule mult_fraction_cong)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   415
  also
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   416
  assume "b \<noteq> 0"  "d \<noteq> 0"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   417
  hence "fract a b * fract c d = fract (a * c) (b * d)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   418
    by (simp add: mult_fraction_def)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   419
  finally show ?thesis by (unfold Fract_def)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   420
qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   421
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   422
theorem inverse_rat: "Fract a b \<noteq> 0 ==> b \<noteq> 0 ==>
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   423
  inverse (Fract a b) = Fract b a"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   424
proof -
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   425
  assume neq: "b \<noteq> 0" and nonzero: "Fract a b \<noteq> 0"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   426
  hence "\<lfloor>fract a b\<rfloor> \<noteq> \<lfloor>0\<rfloor>"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   427
    by (simp add: zero_rat eq_rat is_zero_fraction_iff)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   428
  with _ inverse_fraction_cong [THEN rat_of]
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   429
  have "inverse (Fract a b) = rat_of (inverse (fract a b))"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   430
  proof (rule rat_cond_function)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   431
    fix q assume cond: "\<lfloor>fraction_of q\<rfloor> \<noteq> \<lfloor>0\<rfloor>"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   432
    have "q \<noteq> 0"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   433
    proof (cases q)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   434
      fix a b assume "b \<noteq> 0" and "q = Fract a b"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   435
      from this cond show ?thesis
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   436
        by (simp add: Fract_inverse is_zero_fraction_iff zero_rat eq_rat)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   437
    qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   438
    thus "inverse q == rat_of (inverse (fraction_of q))"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   439
      by (simp add: inverse_rat_def)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   440
  qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   441
  also from neq nonzero have "inverse (fract a b) = fract b a"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   442
    by (simp add: inverse_fraction_def)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   443
  finally show ?thesis by (unfold Fract_def)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   444
qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   445
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   446
theorem divide_rat: "Fract c d \<noteq> 0 ==> b \<noteq> 0 ==> d \<noteq> 0 ==>
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   447
  Fract a b / Fract c d = Fract (a * d) (b * c)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   448
proof -
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   449
  assume neq: "b \<noteq> 0"  "d \<noteq> 0" and nonzero: "Fract c d \<noteq> 0"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   450
  hence "c \<noteq> 0" by (simp add: zero_rat eq_rat)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   451
  with neq nonzero show ?thesis
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   452
    by (simp add: divide_rat_def inverse_rat mult_rat)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   453
qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   454
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   455
theorem le_rat: "b \<noteq> 0 ==> d \<noteq> 0 ==>
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   456
  (Fract a b \<le> Fract c d) = ((a * d) * (b * d) \<le> (c * b) * (b * d))"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   457
proof -
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   458
  have "(Fract a b \<le> Fract c d) = (fract a b \<le> fract c d)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   459
    by (rule rat_function, rule le_rat_def, rule le_fraction_cong)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   460
  also
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   461
  assume "b \<noteq> 0"  "d \<noteq> 0"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   462
  hence "(fract a b \<le> fract c d) = ((a * d) * (b * d) \<le> (c * b) * (b * d))"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   463
    by (simp add: le_fraction_def)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   464
  finally show ?thesis .
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   465
qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   466
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   467
theorem less_rat: "b \<noteq> 0 ==> d \<noteq> 0 ==>
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   468
    (Fract a b < Fract c d) = ((a * d) * (b * d) < (c * b) * (b * d))"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14365
diff changeset
   469
  by (simp add: less_rat_def le_rat eq_rat order_less_le)
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   470
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   471
theorem abs_rat: "b \<noteq> 0 ==> \<bar>Fract a b\<bar> = Fract \<bar>a\<bar> \<bar>b\<bar>"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   472
  by (simp add: abs_rat_def minus_rat zero_rat less_rat eq_rat)
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14365
diff changeset
   473
     (auto simp add: mult_less_0_iff zero_less_mult_iff order_le_less 
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   474
                split: abs_split)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   475
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   476
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   477
subsubsection {* The ordered field of rational numbers *}
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   478
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   479
lemma rat_add_assoc: "(q + r) + s = q + (r + (s::rat))"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   480
  by (induct q, induct r, induct s) 
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   481
     (simp add: add_rat add_ac mult_ac int_distrib)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   482
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   483
lemma rat_add_0: "0 + q = (q::rat)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   484
  by (induct q) (simp add: zero_rat add_rat)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   485
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   486
lemma rat_left_minus: "(-q) + q = (0::rat)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   487
  by (induct q) (simp add: zero_rat minus_rat add_rat eq_rat)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   488
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   489
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   490
instance rat :: field
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   491
proof
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   492
  fix q r s :: rat
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   493
  show "(q + r) + s = q + (r + s)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   494
    by (rule rat_add_assoc)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   495
  show "q + r = r + q"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   496
    by (induct q, induct r) (simp add: add_rat add_ac mult_ac)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   497
  show "0 + q = q"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   498
    by (induct q) (simp add: zero_rat add_rat)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   499
  show "(-q) + q = 0"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   500
    by (rule rat_left_minus)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   501
  show "q - r = q + (-r)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   502
    by (induct q, induct r) (simp add: add_rat minus_rat diff_rat)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   503
  show "(q * r) * s = q * (r * s)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   504
    by (induct q, induct r, induct s) (simp add: mult_rat mult_ac)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   505
  show "q * r = r * q"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   506
    by (induct q, induct r) (simp add: mult_rat mult_ac)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   507
  show "1 * q = q"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   508
    by (induct q) (simp add: one_rat mult_rat)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   509
  show "(q + r) * s = q * s + r * s"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   510
    by (induct q, induct r, induct s) 
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   511
       (simp add: add_rat mult_rat eq_rat int_distrib)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   512
  show "q \<noteq> 0 ==> inverse q * q = 1"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   513
    by (induct q) (simp add: inverse_rat mult_rat one_rat zero_rat eq_rat)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   514
  show "r \<noteq> 0 ==> q / r = q * inverse r"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   515
    by (induct q, induct r)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   516
       (simp add: mult_rat divide_rat inverse_rat zero_rat eq_rat)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   517
  show "0 \<noteq> (1::rat)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   518
    by (simp add: zero_rat one_rat eq_rat) 
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   519
  assume eq: "s+q = s+r" 
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   520
    hence "(-s + s) + q = (-s + s) + r" by (simp only: eq rat_add_assoc)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   521
    thus "q = r" by (simp add: rat_left_minus rat_add_0)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   522
qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   523
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   524
instance rat :: linorder
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   525
proof
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   526
  fix q r s :: rat
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   527
  {
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   528
    assume "q \<le> r" and "r \<le> s"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   529
    show "q \<le> s"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   530
    proof (insert prems, induct q, induct r, induct s)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   531
      fix a b c d e f :: int
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   532
      assume neq: "b \<noteq> 0"  "d \<noteq> 0"  "f \<noteq> 0"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   533
      assume 1: "Fract a b \<le> Fract c d" and 2: "Fract c d \<le> Fract e f"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   534
      show "Fract a b \<le> Fract e f"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   535
      proof -
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   536
        from neq obtain bb: "0 < b * b" and dd: "0 < d * d" and ff: "0 < f * f"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   537
          by (auto simp add: zero_less_mult_iff linorder_neq_iff)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   538
        have "(a * d) * (b * d) * (f * f) \<le> (c * b) * (b * d) * (f * f)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   539
        proof -
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   540
          from neq 1 have "(a * d) * (b * d) \<le> (c * b) * (b * d)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   541
            by (simp add: le_rat)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   542
          with ff show ?thesis by (simp add: mult_le_cancel_right)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   543
        qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   544
        also have "... = (c * f) * (d * f) * (b * b)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   545
          by (simp only: mult_ac)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   546
        also have "... \<le> (e * d) * (d * f) * (b * b)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   547
        proof -
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   548
          from neq 2 have "(c * f) * (d * f) \<le> (e * d) * (d * f)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   549
            by (simp add: le_rat)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   550
          with bb show ?thesis by (simp add: mult_le_cancel_right)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   551
        qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   552
        finally have "(a * f) * (b * f) * (d * d) \<le> e * b * (b * f) * (d * d)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   553
          by (simp only: mult_ac)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   554
        with dd have "(a * f) * (b * f) \<le> (e * b) * (b * f)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   555
          by (simp add: mult_le_cancel_right)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   556
        with neq show ?thesis by (simp add: le_rat)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   557
      qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   558
    qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   559
  next
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   560
    assume "q \<le> r" and "r \<le> q"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   561
    show "q = r"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   562
    proof (insert prems, induct q, induct r)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   563
      fix a b c d :: int
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   564
      assume neq: "b \<noteq> 0"  "d \<noteq> 0"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   565
      assume 1: "Fract a b \<le> Fract c d" and 2: "Fract c d \<le> Fract a b"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   566
      show "Fract a b = Fract c d"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   567
      proof -
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   568
        from neq 1 have "(a * d) * (b * d) \<le> (c * b) * (b * d)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   569
          by (simp add: le_rat)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   570
        also have "... \<le> (a * d) * (b * d)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   571
        proof -
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   572
          from neq 2 have "(c * b) * (d * b) \<le> (a * d) * (d * b)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   573
            by (simp add: le_rat)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   574
          thus ?thesis by (simp only: mult_ac)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   575
        qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   576
        finally have "(a * d) * (b * d) = (c * b) * (b * d)" .
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   577
        moreover from neq have "b * d \<noteq> 0" by simp
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   578
        ultimately have "a * d = c * b" by simp
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   579
        with neq show ?thesis by (simp add: eq_rat)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   580
      qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   581
    qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   582
  next
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   583
    show "q \<le> q"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   584
      by (induct q) (simp add: le_rat)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   585
    show "(q < r) = (q \<le> r \<and> q \<noteq> r)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   586
      by (simp only: less_rat_def)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   587
    show "q \<le> r \<or> r \<le> q"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   588
      by (induct q, induct r) (simp add: le_rat mult_ac, arith)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   589
  }
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   590
qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   591
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   592
instance rat :: ordered_field
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   593
proof
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   594
  fix q r s :: rat
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   595
  show "0 < (1::rat)" 
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   596
    by (simp add: zero_rat one_rat less_rat) 
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   597
  show "q \<le> r ==> s + q \<le> s + r"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   598
  proof (induct q, induct r, induct s)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   599
    fix a b c d e f :: int
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   600
    assume neq: "b \<noteq> 0"  "d \<noteq> 0"  "f \<noteq> 0"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   601
    assume le: "Fract a b \<le> Fract c d"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   602
    show "Fract e f + Fract a b \<le> Fract e f + Fract c d"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   603
    proof -
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   604
      let ?F = "f * f" from neq have F: "0 < ?F"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   605
        by (auto simp add: zero_less_mult_iff)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   606
      from neq le have "(a * d) * (b * d) \<le> (c * b) * (b * d)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   607
        by (simp add: le_rat)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   608
      with F have "(a * d) * (b * d) * ?F * ?F \<le> (c * b) * (b * d) * ?F * ?F"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   609
        by (simp add: mult_le_cancel_right)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   610
      with neq show ?thesis by (simp add: add_rat le_rat mult_ac int_distrib)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   611
    qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   612
  qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   613
  show "q < r ==> 0 < s ==> s * q < s * r"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   614
  proof (induct q, induct r, induct s)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   615
    fix a b c d e f :: int
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   616
    assume neq: "b \<noteq> 0"  "d \<noteq> 0"  "f \<noteq> 0"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   617
    assume le: "Fract a b < Fract c d"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   618
    assume gt: "0 < Fract e f"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   619
    show "Fract e f * Fract a b < Fract e f * Fract c d"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   620
    proof -
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   621
      let ?E = "e * f" and ?F = "f * f"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   622
      from neq gt have "0 < ?E"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14365
diff changeset
   623
        by (auto simp add: zero_rat less_rat le_rat order_less_le eq_rat)
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   624
      moreover from neq have "0 < ?F"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   625
        by (auto simp add: zero_less_mult_iff)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   626
      moreover from neq le have "(a * d) * (b * d) < (c * b) * (b * d)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   627
        by (simp add: less_rat)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   628
      ultimately have "(a * d) * (b * d) * ?E * ?F < (c * b) * (b * d) * ?E * ?F"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   629
        by (simp add: mult_less_cancel_right)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   630
      with neq show ?thesis
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   631
        by (simp add: less_rat mult_rat mult_ac)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   632
    qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   633
  qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   634
  show "\<bar>q\<bar> = (if q < 0 then -q else q)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   635
    by (simp only: abs_rat_def)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   636
qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   637
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   638
instance rat :: division_by_zero
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   639
proof
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   640
  fix x :: rat
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   641
  show "inverse 0 = (0::rat)"  by (simp add: inverse_rat_def)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   642
  show "x/0 = 0"   by (simp add: divide_rat_def inverse_rat_def)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   643
qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   644
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   645
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   646
subsection {* Various Other Results *}
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   647
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   648
lemma minus_rat_cancel [simp]: "b \<noteq> 0 ==> Fract (-a) (-b) = Fract a b"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   649
by (simp add: Fract_equality eq_fraction_iff) 
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   650
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   651
theorem Rat_induct_pos [case_names Fract, induct type: rat]:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   652
  assumes step: "!!a b. 0 < b ==> P (Fract a b)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   653
    shows "P q"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   654
proof (cases q)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   655
  have step': "!!a b. b < 0 ==> P (Fract a b)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   656
  proof -
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   657
    fix a::int and b::int
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   658
    assume b: "b < 0"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   659
    hence "0 < -b" by simp
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   660
    hence "P (Fract (-a) (-b))" by (rule step)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   661
    thus "P (Fract a b)" by (simp add: order_less_imp_not_eq [OF b])
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   662
  qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   663
  case (Fract a b)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   664
  thus "P q" by (force simp add: linorder_neq_iff step step')
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   665
qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   666
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   667
lemma zero_less_Fract_iff:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   668
     "0 < b ==> (0 < Fract a b) = (0 < a)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   669
by (simp add: zero_rat less_rat order_less_imp_not_eq2 zero_less_mult_iff) 
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   670
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14365
diff changeset
   671
lemma Fract_add_one: "n \<noteq> 0 ==> Fract (m + n) n = Fract m n + 1"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14365
diff changeset
   672
apply (insert add_rat [of concl: m n 1 1])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14365
diff changeset
   673
apply (simp add: one_rat  [symmetric]) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14365
diff changeset
   674
done
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14365
diff changeset
   675
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14365
diff changeset
   676
lemma Fract_of_nat_eq: "Fract (of_nat k) 1 = of_nat k"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14365
diff changeset
   677
apply (induct k) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14365
diff changeset
   678
apply (simp add: zero_rat) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14365
diff changeset
   679
apply (simp add: Fract_add_one) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14365
diff changeset
   680
done
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14365
diff changeset
   681
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14365
diff changeset
   682
lemma Fract_of_int_eq: "Fract k 1 = of_int k"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14365
diff changeset
   683
proof (cases k rule: int_cases) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14365
diff changeset
   684
  case (nonneg n)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14365
diff changeset
   685
    thus ?thesis by (simp add: int_eq_of_nat Fract_of_nat_eq)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14365
diff changeset
   686
next 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14365
diff changeset
   687
  case (neg n)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14365
diff changeset
   688
    hence "Fract k 1 = - (Fract (of_nat (Suc n)) 1)"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14365
diff changeset
   689
      by (simp only: minus_rat int_eq_of_nat) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14365
diff changeset
   690
    also have "... =  - (of_nat (Suc n))"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14365
diff changeset
   691
      by (simp only: Fract_of_nat_eq)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14365
diff changeset
   692
    finally show ?thesis 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14365
diff changeset
   693
      by (simp add: only: prems int_eq_of_nat of_int_minus of_int_of_nat_eq) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14365
diff changeset
   694
qed 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14365
diff changeset
   695
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14365
diff changeset
   696
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   697
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   698
subsection{*Numerals and Arithmetic*}
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   699
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   700
instance rat :: number ..
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   701
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   702
primrec (*the type constraint is essential!*)
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   703
  number_of_Pls: "number_of bin.Pls = 0"
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   704
  number_of_Min: "number_of bin.Min = - (1::rat)"
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   705
  number_of_BIT: "number_of(w BIT x) = (if x then 1 else 0) +
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   706
	                               (number_of w) + (number_of w)"
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   707
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   708
declare number_of_Pls [simp del]
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   709
        number_of_Min [simp del]
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   710
        number_of_BIT [simp del]
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   711
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   712
instance rat :: number_ring
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   713
proof
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   714
  show "Numeral0 = (0::rat)" by (rule number_of_Pls)
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   715
  show "-1 = - (1::rat)" by (rule number_of_Min)
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   716
  fix w :: bin and x :: bool
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   717
  show "(number_of (w BIT x) :: rat) =
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   718
        (if x then 1 else 0) + number_of w + number_of w"
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   719
    by (rule number_of_BIT)
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   720
qed
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   721
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   722
declare diff_rat_def [symmetric]
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   723
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   724
use "rat_arith.ML"
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   725
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   726
setup rat_arith_setup
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   727
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   728
end