src/HOL/Real/Rational.thy
changeset 18913 57f19fad8c2a
parent 18372 2bffdf62fe7f
child 18982 a2950f748445
     1.1 --- a/src/HOL/Real/Rational.thy	Thu Feb 02 18:04:10 2006 +0100
     1.2 +++ b/src/HOL/Real/Rational.thy	Thu Feb 02 19:57:13 2006 +0100
     1.3 @@ -6,212 +6,120 @@
     1.4  header {* Rational numbers *}
     1.5  
     1.6  theory Rational
     1.7 -imports Quotient
     1.8 +imports Main
     1.9  uses ("rat_arith.ML")
    1.10  begin
    1.11  
    1.12 -subsection {* Fractions *}
    1.13 -
    1.14 -subsubsection {* The type of fractions *}
    1.15 -
    1.16 -typedef fraction = "{(a, b) :: int \<times> int | a b. b \<noteq> 0}"
    1.17 -proof
    1.18 -  show "(0, 1) \<in> ?fraction" by simp
    1.19 -qed
    1.20 -
    1.21 -constdefs
    1.22 -  fract :: "int => int => fraction"
    1.23 -  "fract a b == Abs_fraction (a, b)"
    1.24 -  num :: "fraction => int"
    1.25 -  "num Q == fst (Rep_fraction Q)"
    1.26 -  den :: "fraction => int"
    1.27 -  "den Q == snd (Rep_fraction Q)"
    1.28 -
    1.29 -lemma fract_num [simp]: "b \<noteq> 0 ==> num (fract a b) = a"
    1.30 -  by (simp add: fract_def num_def fraction_def Abs_fraction_inverse)
    1.31 -
    1.32 -lemma fract_den [simp]: "b \<noteq> 0 ==> den (fract a b) = b"
    1.33 -  by (simp add: fract_def den_def fraction_def Abs_fraction_inverse)
    1.34 -
    1.35 -lemma fraction_cases [case_names fract, cases type: fraction]:
    1.36 -  "(!!a b. Q = fract a b ==> b \<noteq> 0 ==> C) ==> C"
    1.37 -proof -
    1.38 -  assume r: "!!a b. Q = fract a b ==> b \<noteq> 0 ==> C"
    1.39 -  obtain a b where "Q = fract a b" and "b \<noteq> 0"
    1.40 -    by (cases Q) (auto simp add: fract_def fraction_def)
    1.41 -  thus C by (rule r)
    1.42 -qed
    1.43 -
    1.44 -lemma fraction_induct [case_names fract, induct type: fraction]:
    1.45 -    "(!!a b. b \<noteq> 0 ==> P (fract a b)) ==> P Q"
    1.46 -  by (cases Q) simp
    1.47 -
    1.48 +subsection {* Rational numbers *}
    1.49  
    1.50  subsubsection {* Equivalence of fractions *}
    1.51  
    1.52 -instance fraction :: eqv ..
    1.53 +constdefs
    1.54 +  fraction :: "(int \<times> int) set"
    1.55 +   "fraction \<equiv> {x. snd x \<noteq> 0}"
    1.56  
    1.57 -defs (overloaded)
    1.58 -  equiv_fraction_def: "Q \<sim> R == num Q * den R = num R * den Q"
    1.59 +  ratrel :: "((int \<times> int) \<times> (int \<times> int)) set"
    1.60 +   "ratrel \<equiv> {(x,y). snd x \<noteq> 0 \<and> snd y \<noteq> 0 \<and> fst x * snd y = fst y * snd x}"
    1.61  
    1.62 -lemma equiv_fraction_iff [iff]:
    1.63 -    "b \<noteq> 0 ==> b' \<noteq> 0 ==> (fract a b \<sim> fract a' b') = (a * b' = a' * b)"
    1.64 -  by (simp add: equiv_fraction_def)
    1.65 +lemma fraction_iff [simp]: "(x \<in> fraction) = (snd x \<noteq> 0)"
    1.66 +by (simp add: fraction_def)
    1.67  
    1.68 -instance fraction :: equiv
    1.69 -proof
    1.70 -  fix Q R S :: fraction
    1.71 -  {
    1.72 -    show "Q \<sim> Q"
    1.73 -    proof (induct Q)
    1.74 -      fix a b :: int
    1.75 -      assume "b \<noteq> 0" and "b \<noteq> 0"
    1.76 -      with refl show "fract a b \<sim> fract a b" ..
    1.77 -    qed
    1.78 -  next
    1.79 -    assume "Q \<sim> R" and "R \<sim> S"
    1.80 -    show "Q \<sim> S"
    1.81 -    proof (insert prems, induct Q, induct R, induct S)
    1.82 -      fix a b a' b' a'' b'' :: int
    1.83 -      assume b: "b \<noteq> 0" and b': "b' \<noteq> 0" and b'': "b'' \<noteq> 0"
    1.84 -      assume "fract a b \<sim> fract a' b'" hence eq1: "a * b' = a' * b" ..
    1.85 -      assume "fract a' b' \<sim> fract a'' b''" hence eq2: "a' * b'' = a'' * b'" ..
    1.86 -      have "a * b'' = a'' * b"
    1.87 -      proof cases
    1.88 -        assume "a' = 0"
    1.89 -        with b' eq1 eq2 have "a = 0 \<and> a'' = 0" by auto
    1.90 -        thus ?thesis by simp
    1.91 -      next
    1.92 -        assume a': "a' \<noteq> 0"
    1.93 -        from eq1 eq2 have "(a * b') * (a' * b'') = (a' * b) * (a'' * b')" by simp
    1.94 -        hence "(a * b'') * (a' * b') = (a'' * b) * (a' * b')" by (simp only: mult_ac)
    1.95 -        with a' b' show ?thesis by simp
    1.96 -      qed
    1.97 -      thus "fract a b \<sim> fract a'' b''" ..
    1.98 -    qed
    1.99 -  next
   1.100 -    show "Q \<sim> R ==> R \<sim> Q"
   1.101 -    proof (induct Q, induct R)
   1.102 -      fix a b a' b' :: int
   1.103 -      assume b: "b \<noteq> 0" and b': "b' \<noteq> 0"
   1.104 -      assume "fract a b \<sim> fract a' b'"
   1.105 -      hence "a * b' = a' * b" ..
   1.106 -      hence "a' * b = a * b'" ..
   1.107 -      thus "fract a' b' \<sim> fract a b" ..
   1.108 -    qed
   1.109 -  }
   1.110 +lemma ratrel_iff [simp]:
   1.111 +  "((x,y) \<in> ratrel) =
   1.112 +   (snd x \<noteq> 0 \<and> snd y \<noteq> 0 \<and> fst x * snd y = fst y * snd x)"
   1.113 +by (simp add: ratrel_def)
   1.114 +
   1.115 +lemma refl_ratrel: "refl fraction ratrel"
   1.116 +by (auto simp add: refl_def fraction_def ratrel_def)
   1.117 +
   1.118 +lemma sym_ratrel: "sym ratrel"
   1.119 +by (simp add: ratrel_def sym_def)
   1.120 +
   1.121 +lemma trans_ratrel_lemma:
   1.122 +  assumes 1: "a * b' = a' * b"
   1.123 +  assumes 2: "a' * b'' = a'' * b'"
   1.124 +  assumes 3: "b' \<noteq> (0::int)"
   1.125 +  shows "a * b'' = a'' * b"
   1.126 +proof -
   1.127 +  have "b' * (a * b'') = b'' * (a * b')" by simp
   1.128 +  also note 1
   1.129 +  also have "b'' * (a' * b) = b * (a' * b'')" by simp
   1.130 +  also note 2
   1.131 +  also have "b * (a'' * b') = b' * (a'' * b)" by simp
   1.132 +  finally have "b' * (a * b'') = b' * (a'' * b)" .
   1.133 +  with 3 show "a * b'' = a'' * b" by simp
   1.134  qed
   1.135  
   1.136 -lemma eq_fraction_iff [iff]:
   1.137 -    "b \<noteq> 0 ==> b' \<noteq> 0 ==> (\<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor>) = (a * b' = a' * b)"
   1.138 -  by (simp add: equiv_fraction_iff quot_equality)
   1.139 +lemma trans_ratrel: "trans ratrel"
   1.140 +by (auto simp add: trans_def elim: trans_ratrel_lemma)
   1.141  
   1.142 +lemma equiv_ratrel: "equiv fraction ratrel"
   1.143 +by (rule equiv.intro [OF refl_ratrel sym_ratrel trans_ratrel])
   1.144  
   1.145 -subsubsection {* Operations on fractions *}
   1.146 +lemmas equiv_ratrel_iff [iff] = eq_equiv_class_iff [OF equiv_ratrel]
   1.147  
   1.148 -text {*
   1.149 - We define the basic arithmetic operations on fractions and
   1.150 - demonstrate their ``well-definedness'', i.e.\ congruence with respect
   1.151 - to equivalence of fractions.
   1.152 -*}
   1.153 +lemma equiv_ratrel_iff2:
   1.154 +  "\<lbrakk>snd x \<noteq> 0; snd y \<noteq> 0\<rbrakk>
   1.155 +    \<Longrightarrow> (ratrel `` {x} = ratrel `` {y}) = ((x,y) \<in> ratrel)"
   1.156 +by (rule eq_equiv_class_iff [OF equiv_ratrel], simp_all)
   1.157  
   1.158 -instance fraction :: "{zero, one, plus, minus, times, inverse, ord}" ..
   1.159  
   1.160 -defs (overloaded)
   1.161 -  zero_fraction_def: "0 == fract 0 1"
   1.162 -  one_fraction_def: "1 == fract 1 1"
   1.163 -  add_fraction_def: "Q + R ==
   1.164 -    fract (num Q * den R + num R * den Q) (den Q * den R)"
   1.165 -  minus_fraction_def: "-Q == fract (-(num Q)) (den Q)"
   1.166 -  mult_fraction_def: "Q * R == fract (num Q * num R) (den Q * den R)"
   1.167 -  inverse_fraction_def: "inverse Q == fract (den Q) (num Q)"
   1.168 -  le_fraction_def: "Q \<le> R ==
   1.169 -    (num Q * den R) * (den Q * den R) \<le> (num R * den Q) * (den Q * den R)"
   1.170 +subsubsection {* The type of rational numbers *}
   1.171  
   1.172 -lemma is_zero_fraction_iff: "b \<noteq> 0 ==> (\<lfloor>fract a b\<rfloor> = \<lfloor>0\<rfloor>) = (a = 0)"
   1.173 -  by (simp add: zero_fraction_def eq_fraction_iff)
   1.174 -
   1.175 -theorem add_fraction_cong:
   1.176 -  "\<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor> ==> \<lfloor>fract c d\<rfloor> = \<lfloor>fract c' d'\<rfloor>
   1.177 -    ==> b \<noteq> 0 ==> b' \<noteq> 0 ==> d \<noteq> 0 ==> d' \<noteq> 0
   1.178 -    ==> \<lfloor>fract a b + fract c d\<rfloor> = \<lfloor>fract a' b' + fract c' d'\<rfloor>"
   1.179 -proof -
   1.180 -  assume neq: "b \<noteq> 0"  "b' \<noteq> 0"  "d \<noteq> 0"  "d' \<noteq> 0"
   1.181 -  assume "\<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor>" hence eq1: "a * b' = a' * b" ..
   1.182 -  assume "\<lfloor>fract c d\<rfloor> = \<lfloor>fract c' d'\<rfloor>" hence eq2: "c * d' = c' * d" ..
   1.183 -  have "\<lfloor>fract (a * d + c * b) (b * d)\<rfloor> = \<lfloor>fract (a' * d' + c' * b') (b' * d')\<rfloor>"
   1.184 -  proof
   1.185 -    show "(a * d + c * b) * (b' * d') = (a' * d' + c' * b') * (b * d)"
   1.186 -      (is "?lhs = ?rhs")
   1.187 -    proof -
   1.188 -      have "?lhs = (a * b') * (d * d') + (c * d') * (b * b')"
   1.189 -        by (simp add: int_distrib mult_ac)
   1.190 -      also have "... = (a' * b) * (d * d') + (c' * d) * (b * b')"
   1.191 -        by (simp only: eq1 eq2)
   1.192 -      also have "... = ?rhs"
   1.193 -        by (simp add: int_distrib mult_ac)
   1.194 -      finally show "?lhs = ?rhs" .
   1.195 -    qed
   1.196 -    from neq show "b * d \<noteq> 0" by simp
   1.197 -    from neq show "b' * d' \<noteq> 0" by simp
   1.198 -  qed
   1.199 -  with neq show ?thesis by (simp add: add_fraction_def)
   1.200 +typedef (Rat) rat = "fraction//ratrel"
   1.201 +proof
   1.202 +  have "(0,1) \<in> fraction" by (simp add: fraction_def)
   1.203 +  thus "ratrel``{(0,1)} \<in> fraction//ratrel" by (rule quotientI)
   1.204  qed
   1.205  
   1.206 -theorem minus_fraction_cong:
   1.207 -  "\<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor> ==> b \<noteq> 0 ==> b' \<noteq> 0
   1.208 -    ==> \<lfloor>-(fract a b)\<rfloor> = \<lfloor>-(fract a' b')\<rfloor>"
   1.209 -proof -
   1.210 -  assume neq: "b \<noteq> 0"  "b' \<noteq> 0"
   1.211 -  assume "\<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor>"
   1.212 -  hence "a * b' = a' * b" ..
   1.213 -  hence "-a * b' = -a' * b" by simp
   1.214 -  hence "\<lfloor>fract (-a) b\<rfloor> = \<lfloor>fract (-a') b'\<rfloor>" ..
   1.215 -  with neq show ?thesis by (simp add: minus_fraction_def)
   1.216 -qed
   1.217 +lemma ratrel_in_Rat [simp]: "snd x \<noteq> 0 \<Longrightarrow> ratrel``{x} \<in> Rat"
   1.218 +by (simp add: Rat_def quotientI)
   1.219  
   1.220 -theorem mult_fraction_cong:
   1.221 -  "\<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor> ==> \<lfloor>fract c d\<rfloor> = \<lfloor>fract c' d'\<rfloor>
   1.222 -    ==> b \<noteq> 0 ==> b' \<noteq> 0 ==> d \<noteq> 0 ==> d' \<noteq> 0
   1.223 -    ==> \<lfloor>fract a b * fract c d\<rfloor> = \<lfloor>fract a' b' * fract c' d'\<rfloor>"
   1.224 -proof -
   1.225 +declare Abs_Rat_inject [simp]  Abs_Rat_inverse [simp]
   1.226 +
   1.227 +
   1.228 +constdefs
   1.229 +  Fract :: "int \<Rightarrow> int \<Rightarrow> rat"
   1.230 +  "Fract a b \<equiv> Abs_Rat (ratrel``{(a,b)})"
   1.231 +
   1.232 +theorem Rat_cases [case_names Fract, cases type: rat]:
   1.233 +  "(!!a b. q = Fract a b ==> b \<noteq> 0 ==> C) ==> C"
   1.234 +by (cases q, clarsimp simp add: Fract_def Rat_def fraction_def quotient_def)
   1.235 +
   1.236 +theorem Rat_induct [case_names Fract, induct type: rat]:
   1.237 +    "(!!a b. b \<noteq> 0 ==> P (Fract a b)) ==> P q"
   1.238 +  by (cases q) simp
   1.239 +
   1.240 +
   1.241 +subsubsection {* Congruence lemmas *}
   1.242 +
   1.243 +lemma add_congruent2:
   1.244 +     "(\<lambda>x y. ratrel``{(fst x * snd y + fst y * snd x, snd x * snd y)})
   1.245 +      respects2 ratrel"
   1.246 +apply (rule equiv_ratrel [THEN congruent2_commuteI])
   1.247 +apply (simp_all add: left_distrib)
   1.248 +done
   1.249 +
   1.250 +lemma minus_congruent:
   1.251 +  "(\<lambda>x. ratrel``{(- fst x, snd x)}) respects ratrel"
   1.252 +by (simp add: congruent_def)
   1.253 +
   1.254 +lemma mult_congruent2:
   1.255 +  "(\<lambda>x y. ratrel``{(fst x * fst y, snd x * snd y)}) respects2 ratrel"
   1.256 +by (rule equiv_ratrel [THEN congruent2_commuteI], simp_all)
   1.257 +
   1.258 +lemma inverse_congruent:
   1.259 +  "(\<lambda>x. ratrel``{if fst x=0 then (0,1) else (snd x, fst x)}) respects ratrel"
   1.260 +by (auto simp add: congruent_def mult_commute)
   1.261 +
   1.262 +lemma le_congruent2:
   1.263 +  "(\<lambda>x y. (fst x * snd y)*(snd x * snd y) \<le> (fst y * snd x)*(snd x * snd y))
   1.264 +   respects2 ratrel"
   1.265 +proof (clarsimp simp add: congruent2_def)
   1.266 +  fix a b a' b' c d c' d'::int
   1.267    assume neq: "b \<noteq> 0"  "b' \<noteq> 0"  "d \<noteq> 0"  "d' \<noteq> 0"
   1.268 -  assume "\<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor>" hence eq1: "a * b' = a' * b" ..
   1.269 -  assume "\<lfloor>fract c d\<rfloor> = \<lfloor>fract c' d'\<rfloor>" hence eq2: "c * d' = c' * d" ..
   1.270 -  have "\<lfloor>fract (a * c) (b * d)\<rfloor> = \<lfloor>fract (a' * c') (b' * d')\<rfloor>"
   1.271 -  proof
   1.272 -    from eq1 eq2 have "(a * b') * (c * d') = (a' * b) * (c' * d)" by simp
   1.273 -    thus "(a * c) * (b' * d') = (a' * c') * (b * d)" by (simp add: mult_ac)
   1.274 -    from neq show "b * d \<noteq> 0" by simp
   1.275 -    from neq show "b' * d' \<noteq> 0" by simp
   1.276 -  qed
   1.277 -  with neq show "\<lfloor>fract a b * fract c d\<rfloor> = \<lfloor>fract a' b' * fract c' d'\<rfloor>"
   1.278 -    by (simp add: mult_fraction_def)
   1.279 -qed
   1.280 -
   1.281 -theorem inverse_fraction_cong:
   1.282 -  "\<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>
   1.283 -    ==> b \<noteq> 0 ==> b' \<noteq> 0
   1.284 -    ==> \<lfloor>inverse (fract a b)\<rfloor> = \<lfloor>inverse (fract a' b')\<rfloor>"
   1.285 -proof -
   1.286 -  assume neq: "b \<noteq> 0"  "b' \<noteq> 0"
   1.287 -  assume "\<lfloor>fract a b\<rfloor> \<noteq> \<lfloor>0\<rfloor>" and "\<lfloor>fract a' b'\<rfloor> \<noteq> \<lfloor>0\<rfloor>"
   1.288 -  with neq obtain "a \<noteq> 0" and "a' \<noteq> 0" by (simp add: is_zero_fraction_iff)
   1.289 -  assume "\<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor>"
   1.290 -  hence "a * b' = a' * b" ..
   1.291 -  hence "b * a' = b' * a" by (simp only: mult_ac)
   1.292 -  hence "\<lfloor>fract b a\<rfloor> = \<lfloor>fract b' a'\<rfloor>" ..
   1.293 -  with neq show ?thesis by (simp add: inverse_fraction_def)
   1.294 -qed
   1.295 -
   1.296 -theorem le_fraction_cong:
   1.297 -  "\<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor> ==> \<lfloor>fract c d\<rfloor> = \<lfloor>fract c' d'\<rfloor>
   1.298 -    ==> b \<noteq> 0 ==> b' \<noteq> 0 ==> d \<noteq> 0 ==> d' \<noteq> 0
   1.299 -    ==> (fract a b \<le> fract c d) = (fract a' b' \<le> fract c' d')"
   1.300 -proof -
   1.301 -  assume neq: "b \<noteq> 0"  "b' \<noteq> 0"  "d \<noteq> 0"  "d' \<noteq> 0"
   1.302 -  assume "\<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor>" hence eq1: "a * b' = a' * b" ..
   1.303 -  assume "\<lfloor>fract c d\<rfloor> = \<lfloor>fract c' d'\<rfloor>" hence eq2: "c * d' = c' * d" ..
   1.304 +  assume eq1: "a * b' = a' * b"
   1.305 +  assume eq2: "c * d' = c' * d"
   1.306  
   1.307    let ?le = "\<lambda>a b c d. ((a * d) * (b * d) \<le> (c * b) * (b * d))"
   1.308    {
   1.309 @@ -241,215 +149,120 @@
   1.310      by (simp add: mult_ac)
   1.311    also from D have "... = ?le a' b' c' d'"
   1.312      by (rule le_factor [symmetric])
   1.313 -  finally have "?le a b c d = ?le a' b' c' d'" .
   1.314 -  with neq show ?thesis by (simp add: le_fraction_def)
   1.315 +  finally show "?le a b c d = ?le a' b' c' d'" .
   1.316  qed
   1.317  
   1.318 +lemma All_equiv_class:
   1.319 +  "equiv A r ==> f respects r ==> a \<in> A
   1.320 +    ==> (\<forall>x \<in> r``{a}. f x) = f a"
   1.321 +apply safe
   1.322 +apply (drule (1) equiv_class_self)
   1.323 +apply simp
   1.324 +apply (drule (1) congruent.congruent)
   1.325 +apply simp
   1.326 +done
   1.327  
   1.328 -subsection {* Rational numbers *}
   1.329 +lemma congruent2_implies_congruent_All:
   1.330 +  "equiv A1 r1 ==> equiv A2 r2 ==> congruent2 r1 r2 f ==> a \<in> A2 ==>
   1.331 +    congruent r1 (\<lambda>x1. \<forall>x2 \<in> r2``{a}. f x1 x2)"
   1.332 +  apply (unfold congruent_def)
   1.333 +  apply clarify
   1.334 +  apply (rule equiv_type [THEN subsetD, THEN SigmaE2], assumption+)
   1.335 +  apply (simp add: UN_equiv_class congruent2_implies_congruent)
   1.336 +  apply (unfold congruent2_def equiv_def refl_def)
   1.337 +  apply (blast del: equalityI)
   1.338 +  done
   1.339  
   1.340 -subsubsection {* The type of rational numbers *}
   1.341 +lemma All_equiv_class2:
   1.342 +  "equiv A1 r1 ==> equiv A2 r2 ==> congruent2 r1 r2 f ==> a1 \<in> A1 ==> a2 \<in> A2
   1.343 +    ==> (\<forall>x1 \<in> r1``{a1}. \<forall>x2 \<in> r2``{a2}. f x1 x2) = f a1 a2"
   1.344 +  by (simp add: All_equiv_class congruent2_implies_congruent
   1.345 +    congruent2_implies_congruent_All)
   1.346  
   1.347 -typedef (Rat)
   1.348 -  rat = "UNIV :: fraction quot set" ..
   1.349 -
   1.350 -lemma RatI [intro, simp]: "Q \<in> Rat"
   1.351 -  by (simp add: Rat_def)
   1.352 -
   1.353 -constdefs
   1.354 -  fraction_of :: "rat => fraction"
   1.355 -  "fraction_of q == pick (Rep_Rat q)"
   1.356 -  rat_of :: "fraction => rat"
   1.357 -  "rat_of Q == Abs_Rat \<lfloor>Q\<rfloor>"
   1.358 -
   1.359 -theorem rat_of_equality [iff?]: "(rat_of Q = rat_of Q') = (\<lfloor>Q\<rfloor> = \<lfloor>Q'\<rfloor>)"
   1.360 -  by (simp add: rat_of_def Abs_Rat_inject)
   1.361 -
   1.362 -lemma rat_of: "\<lfloor>Q\<rfloor> = \<lfloor>Q'\<rfloor> ==> rat_of Q = rat_of Q'" ..
   1.363 -
   1.364 -constdefs
   1.365 -  Fract :: "int => int => rat"
   1.366 -  "Fract a b == rat_of (fract a b)"
   1.367 -
   1.368 -theorem Fract_inverse: "\<lfloor>fraction_of (Fract a b)\<rfloor> = \<lfloor>fract a b\<rfloor>"
   1.369 -  by (simp add: fraction_of_def rat_of_def Fract_def Abs_Rat_inverse pick_inverse)
   1.370 -
   1.371 -theorem Fract_equality [iff?]:
   1.372 -    "(Fract a b = Fract c d) = (\<lfloor>fract a b\<rfloor> = \<lfloor>fract c d\<rfloor>)"
   1.373 -  by (simp add: Fract_def rat_of_equality)
   1.374 -
   1.375 -theorem eq_rat:
   1.376 -    "b \<noteq> 0 ==> d \<noteq> 0 ==> (Fract a b = Fract c d) = (a * d = c * b)"
   1.377 -  by (simp add: Fract_equality eq_fraction_iff)
   1.378 -
   1.379 -theorem Rat_cases [case_names Fract, cases type: rat]:
   1.380 -  "(!!a b. q = Fract a b ==> b \<noteq> 0 ==> C) ==> C"
   1.381 -proof -
   1.382 -  assume r: "!!a b. q = Fract a b ==> b \<noteq> 0 ==> C"
   1.383 -  obtain x where "q = Abs_Rat x" by (cases q)
   1.384 -  moreover obtain Q where "x = \<lfloor>Q\<rfloor>" by (cases x)
   1.385 -  moreover obtain a b where "Q = fract a b" and "b \<noteq> 0" by (cases Q)
   1.386 -  ultimately have "q = Fract a b" by (simp only: Fract_def rat_of_def)
   1.387 -  thus ?thesis by (rule r)
   1.388 -qed
   1.389 -
   1.390 -theorem Rat_induct [case_names Fract, induct type: rat]:
   1.391 -    "(!!a b. b \<noteq> 0 ==> P (Fract a b)) ==> P q"
   1.392 -  by (cases q) simp
   1.393 -
   1.394 -
   1.395 -subsubsection {* Canonical function definitions *}
   1.396 -
   1.397 -text {*
   1.398 -  Note that the unconditional version below is much easier to read.
   1.399 -*}
   1.400 -
   1.401 -theorem rat_cond_function:
   1.402 -  "(!!q r. P \<lfloor>fraction_of q\<rfloor> \<lfloor>fraction_of r\<rfloor> ==>
   1.403 -      f q r == g (fraction_of q) (fraction_of r)) ==>
   1.404 -    (!!a b a' b' c d c' d'.
   1.405 -      \<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor> ==> \<lfloor>fract c d\<rfloor> = \<lfloor>fract c' d'\<rfloor> ==>
   1.406 -      P \<lfloor>fract a b\<rfloor> \<lfloor>fract c d\<rfloor> ==> P \<lfloor>fract a' b'\<rfloor> \<lfloor>fract c' d'\<rfloor> ==>
   1.407 -      b \<noteq> 0 ==> b' \<noteq> 0 ==> d \<noteq> 0 ==> d' \<noteq> 0 ==>
   1.408 -      g (fract a b) (fract c d) = g (fract a' b') (fract c' d')) ==>
   1.409 -    P \<lfloor>fract a b\<rfloor> \<lfloor>fract c d\<rfloor> ==>
   1.410 -      f (Fract a b) (Fract c d) = g (fract a b) (fract c d)"
   1.411 -  (is "PROP ?eq ==> PROP ?cong ==> ?P ==> _")
   1.412 -proof -
   1.413 -  assume eq: "PROP ?eq" and cong: "PROP ?cong" and P: ?P
   1.414 -  have "f (Abs_Rat \<lfloor>fract a b\<rfloor>) (Abs_Rat \<lfloor>fract c d\<rfloor>) = g (fract a b) (fract c d)"
   1.415 -  proof (rule quot_cond_function)
   1.416 -    fix X Y assume "P X Y"
   1.417 -    with eq show "f (Abs_Rat X) (Abs_Rat Y) == g (pick X) (pick Y)"
   1.418 -      by (simp add: fraction_of_def pick_inverse Abs_Rat_inverse)
   1.419 -  next
   1.420 -    fix Q Q' R R' :: fraction
   1.421 -    show "\<lfloor>Q\<rfloor> = \<lfloor>Q'\<rfloor> ==> \<lfloor>R\<rfloor> = \<lfloor>R'\<rfloor> ==>
   1.422 -        P \<lfloor>Q\<rfloor> \<lfloor>R\<rfloor> ==> P \<lfloor>Q'\<rfloor> \<lfloor>R'\<rfloor> ==> g Q R = g Q' R'"
   1.423 -      by (induct Q, induct Q', induct R, induct R') (rule cong)
   1.424 -  qed
   1.425 -  thus ?thesis by (unfold Fract_def rat_of_def)
   1.426 -qed
   1.427 -
   1.428 -theorem rat_function:
   1.429 -  assumes "!!q r. f q r == g (fraction_of q) (fraction_of r)"
   1.430 -    and "!!a b a' b' c d c' d'.
   1.431 -      \<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor> ==> \<lfloor>fract c d\<rfloor> = \<lfloor>fract c' d'\<rfloor> ==>
   1.432 -      b \<noteq> 0 ==> b' \<noteq> 0 ==> d \<noteq> 0 ==> d' \<noteq> 0 ==>
   1.433 -      g (fract a b) (fract c d) = g (fract a' b') (fract c' d')"
   1.434 -  shows "f (Fract a b) (Fract c d) = g (fract a b) (fract c d)"
   1.435 -  using prems and TrueI by (rule rat_cond_function)
   1.436 +lemmas UN_ratrel = UN_equiv_class [OF equiv_ratrel]
   1.437 +lemmas UN_ratrel2 = UN_equiv_class2 [OF equiv_ratrel equiv_ratrel]
   1.438 +lemmas All_ratrel2 = All_equiv_class2 [OF equiv_ratrel equiv_ratrel]
   1.439  
   1.440  
   1.441  subsubsection {* Standard operations on rational numbers *}
   1.442  
   1.443 -instance rat :: "{zero, one, plus, minus, times, inverse, ord}" ..
   1.444 +instance rat :: "{ord, zero, one, plus, times, minus, inverse}" ..
   1.445  
   1.446  defs (overloaded)
   1.447 -  zero_rat_def: "0 == rat_of 0"
   1.448 -  one_rat_def: "1 == rat_of 1"
   1.449 -  add_rat_def: "q + r == rat_of (fraction_of q + fraction_of r)"
   1.450 -  minus_rat_def: "-q == rat_of (-(fraction_of q))"
   1.451 -  diff_rat_def: "q - r == q + (-(r::rat))"
   1.452 -  mult_rat_def: "q * r == rat_of (fraction_of q * fraction_of r)"
   1.453 -  inverse_rat_def: "inverse q ==
   1.454 -                    if q=0 then 0 else rat_of (inverse (fraction_of q))"
   1.455 -  divide_rat_def: "q / r == q * inverse (r::rat)"
   1.456 -  le_rat_def: "q \<le> r == fraction_of q \<le> fraction_of r"
   1.457 -  less_rat_def: "q < r == q \<le> r \<and> q \<noteq> (r::rat)"
   1.458 +  Zero_rat_def:  "0 == Fract 0 1"
   1.459 +  One_rat_def:   "1 == Fract 1 1"
   1.460 +
   1.461 +  add_rat_def:
   1.462 +   "q + r ==
   1.463 +       Abs_Rat (\<Union>x \<in> Rep_Rat q. \<Union>y \<in> Rep_Rat r.
   1.464 +           ratrel``{(fst x * snd y + fst y * snd x, snd x * snd y)})"
   1.465 +
   1.466 +  minus_rat_def:
   1.467 +    "- q == Abs_Rat (\<Union>x \<in> Rep_Rat q. ratrel``{(- fst x, snd x)})"
   1.468 +
   1.469 +  diff_rat_def:  "q - r == q + - (r::rat)"
   1.470 +
   1.471 +  mult_rat_def:
   1.472 +   "q * r ==
   1.473 +       Abs_Rat (\<Union>x \<in> Rep_Rat q. \<Union>y \<in> Rep_Rat r.
   1.474 +           ratrel``{(fst x * fst y, snd x * snd y)})"
   1.475 +
   1.476 +  inverse_rat_def:
   1.477 +    "inverse q ==
   1.478 +        Abs_Rat (\<Union>x \<in> Rep_Rat q.
   1.479 +            ratrel``{if fst x=0 then (0,1) else (snd x, fst x)})"
   1.480 +
   1.481 +  divide_rat_def:  "q / r == q * inverse (r::rat)"
   1.482 +
   1.483 +  le_rat_def:
   1.484 +   "q \<le> (r::rat) ==
   1.485 +    \<forall>x \<in> Rep_Rat q. \<forall>y \<in> Rep_Rat r.
   1.486 +        (fst x * snd y)*(snd x * snd y) \<le> (fst y * snd x)*(snd x * snd y)"
   1.487 +
   1.488 +  less_rat_def: "(z < (w::rat)) == (z \<le> w & z \<noteq> w)"
   1.489 +
   1.490    abs_rat_def: "\<bar>q\<bar> == if q < 0 then -q else (q::rat)"
   1.491  
   1.492 -theorem zero_rat: "0 = Fract 0 1"
   1.493 -  by (simp add: zero_rat_def zero_fraction_def rat_of_def Fract_def)
   1.494 +lemma zero_rat: "0 = Fract 0 1"
   1.495 +by (simp add: Zero_rat_def)
   1.496  
   1.497 -theorem one_rat: "1 = Fract 1 1"
   1.498 -  by (simp add: one_rat_def one_fraction_def rat_of_def Fract_def)
   1.499 +lemma one_rat: "1 = Fract 1 1"
   1.500 +by (simp add: One_rat_def)
   1.501 +
   1.502 +theorem eq_rat: "b \<noteq> 0 ==> d \<noteq> 0 ==>
   1.503 +  (Fract a b = Fract c d) = (a * d = c * b)"
   1.504 +by (simp add: Fract_def)
   1.505  
   1.506  theorem add_rat: "b \<noteq> 0 ==> d \<noteq> 0 ==>
   1.507    Fract a b + Fract c d = Fract (a * d + c * b) (b * d)"
   1.508 -proof -
   1.509 -  have "Fract a b + Fract c d = rat_of (fract a b + fract c d)"
   1.510 -    by (rule rat_function, rule add_rat_def, rule rat_of, rule add_fraction_cong)
   1.511 -  also
   1.512 -  assume "b \<noteq> 0"  "d \<noteq> 0"
   1.513 -  hence "fract a b + fract c d = fract (a * d + c * b) (b * d)"
   1.514 -    by (simp add: add_fraction_def)
   1.515 -  finally show ?thesis by (unfold Fract_def)
   1.516 -qed
   1.517 +by (simp add: Fract_def add_rat_def add_congruent2 UN_ratrel2)
   1.518  
   1.519  theorem minus_rat: "b \<noteq> 0 ==> -(Fract a b) = Fract (-a) b"
   1.520 -proof -
   1.521 -  have "-(Fract a b) = rat_of (-(fract a b))"
   1.522 -    by (rule rat_function, rule minus_rat_def, rule rat_of, rule minus_fraction_cong)
   1.523 -  also assume "b \<noteq> 0" hence "-(fract a b) = fract (-a) b"
   1.524 -    by (simp add: minus_fraction_def)
   1.525 -  finally show ?thesis by (unfold Fract_def)
   1.526 -qed
   1.527 +by (simp add: Fract_def minus_rat_def minus_congruent UN_ratrel)
   1.528  
   1.529  theorem diff_rat: "b \<noteq> 0 ==> d \<noteq> 0 ==>
   1.530      Fract a b - Fract c d = Fract (a * d - c * b) (b * d)"
   1.531 -  by (simp add: diff_rat_def add_rat minus_rat)
   1.532 +by (simp add: diff_rat_def add_rat minus_rat)
   1.533  
   1.534  theorem mult_rat: "b \<noteq> 0 ==> d \<noteq> 0 ==>
   1.535    Fract a b * Fract c d = Fract (a * c) (b * d)"
   1.536 -proof -
   1.537 -  have "Fract a b * Fract c d = rat_of (fract a b * fract c d)"
   1.538 -    by (rule rat_function, rule mult_rat_def, rule rat_of, rule mult_fraction_cong)
   1.539 -  also
   1.540 -  assume "b \<noteq> 0"  "d \<noteq> 0"
   1.541 -  hence "fract a b * fract c d = fract (a * c) (b * d)"
   1.542 -    by (simp add: mult_fraction_def)
   1.543 -  finally show ?thesis by (unfold Fract_def)
   1.544 -qed
   1.545 +by (simp add: Fract_def mult_rat_def mult_congruent2 UN_ratrel2)
   1.546  
   1.547 -theorem inverse_rat: "Fract a b \<noteq> 0 ==> b \<noteq> 0 ==>
   1.548 +theorem inverse_rat: "a \<noteq> 0 ==> b \<noteq> 0 ==>
   1.549    inverse (Fract a b) = Fract b a"
   1.550 -proof -
   1.551 -  assume neq: "b \<noteq> 0" and nonzero: "Fract a b \<noteq> 0"
   1.552 -  hence "\<lfloor>fract a b\<rfloor> \<noteq> \<lfloor>0\<rfloor>"
   1.553 -    by (simp add: zero_rat eq_rat is_zero_fraction_iff)
   1.554 -  with _ inverse_fraction_cong [THEN rat_of]
   1.555 -  have "inverse (Fract a b) = rat_of (inverse (fract a b))"
   1.556 -  proof (rule rat_cond_function)
   1.557 -    fix q assume cond: "\<lfloor>fraction_of q\<rfloor> \<noteq> \<lfloor>0\<rfloor>"
   1.558 -    have "q \<noteq> 0"
   1.559 -    proof (cases q)
   1.560 -      fix a b assume "b \<noteq> 0" and "q = Fract a b"
   1.561 -      from this cond show ?thesis
   1.562 -        by (simp add: Fract_inverse is_zero_fraction_iff zero_rat eq_rat)
   1.563 -    qed
   1.564 -    thus "inverse q == rat_of (inverse (fraction_of q))"
   1.565 -      by (simp add: inverse_rat_def)
   1.566 -  qed
   1.567 -  also from neq nonzero have "inverse (fract a b) = fract b a"
   1.568 -    by (simp add: inverse_fraction_def)
   1.569 -  finally show ?thesis by (unfold Fract_def)
   1.570 -qed
   1.571 +by (simp add: Fract_def inverse_rat_def inverse_congruent UN_ratrel)
   1.572  
   1.573 -theorem divide_rat: "Fract c d \<noteq> 0 ==> b \<noteq> 0 ==> d \<noteq> 0 ==>
   1.574 +theorem divide_rat: "c \<noteq> 0 ==> b \<noteq> 0 ==> d \<noteq> 0 ==>
   1.575    Fract a b / Fract c d = Fract (a * d) (b * c)"
   1.576 -proof -
   1.577 -  assume neq: "b \<noteq> 0"  "d \<noteq> 0" and nonzero: "Fract c d \<noteq> 0"
   1.578 -  hence "c \<noteq> 0" by (simp add: zero_rat eq_rat)
   1.579 -  with neq nonzero show ?thesis
   1.580 -    by (simp add: divide_rat_def inverse_rat mult_rat)
   1.581 -qed
   1.582 +by (simp add: divide_rat_def inverse_rat mult_rat)
   1.583  
   1.584  theorem le_rat: "b \<noteq> 0 ==> d \<noteq> 0 ==>
   1.585    (Fract a b \<le> Fract c d) = ((a * d) * (b * d) \<le> (c * b) * (b * d))"
   1.586 -proof -
   1.587 -  have "(Fract a b \<le> Fract c d) = (fract a b \<le> fract c d)"
   1.588 -    by (rule rat_function, rule le_rat_def, rule le_fraction_cong)
   1.589 -  also
   1.590 -  assume "b \<noteq> 0"  "d \<noteq> 0"
   1.591 -  hence "(fract a b \<le> fract c d) = ((a * d) * (b * d) \<le> (c * b) * (b * d))"
   1.592 -    by (simp add: le_fraction_def)
   1.593 -  finally show ?thesis .
   1.594 -qed
   1.595 +by (simp add: Fract_def le_rat_def le_congruent2 All_ratrel2)
   1.596  
   1.597  theorem less_rat: "b \<noteq> 0 ==> d \<noteq> 0 ==>
   1.598      (Fract a b < Fract c d) = ((a * d) * (b * d) < (c * b) * (b * d))"
   1.599 -  by (simp add: less_rat_def le_rat eq_rat order_less_le)
   1.600 +by (simp add: less_rat_def le_rat eq_rat order_less_le)
   1.601  
   1.602  theorem abs_rat: "b \<noteq> 0 ==> \<bar>Fract a b\<bar> = Fract \<bar>a\<bar> \<bar>b\<bar>"
   1.603    by (simp add: abs_rat_def minus_rat zero_rat less_rat eq_rat)
   1.604 @@ -474,13 +287,14 @@
   1.605  proof
   1.606    fix q r s :: rat
   1.607    show "(q + r) + s = q + (r + s)"
   1.608 -    by (rule rat_add_assoc)
   1.609 +    by (induct q, induct r, induct s)
   1.610 +       (simp add: add_rat add_ac mult_ac int_distrib)
   1.611    show "q + r = r + q"
   1.612      by (induct q, induct r) (simp add: add_rat add_ac mult_ac)
   1.613    show "0 + q = q"
   1.614      by (induct q) (simp add: zero_rat add_rat)
   1.615    show "(-q) + q = 0"
   1.616 -    by (rule rat_left_minus)
   1.617 +    by (induct q) (simp add: zero_rat minus_rat add_rat eq_rat)
   1.618    show "q - r = q + (-r)"
   1.619      by (induct q, induct r) (simp add: add_rat minus_rat diff_rat)
   1.620    show "(q * r) * s = q * (r * s)"
   1.621 @@ -564,7 +378,8 @@
   1.622      show "(q < r) = (q \<le> r \<and> q \<noteq> r)"
   1.623        by (simp only: less_rat_def)
   1.624      show "q \<le> r \<or> r \<le> q"
   1.625 -      by (induct q, induct r) (simp add: le_rat mult_ac, arith)
   1.626 +      by (induct q, induct r)
   1.627 +         (simp add: le_rat mult_commute, rule linorder_linear)
   1.628    }
   1.629  qed
   1.630  
   1.631 @@ -614,14 +429,16 @@
   1.632  
   1.633  instance rat :: division_by_zero
   1.634  proof
   1.635 -  show "inverse 0 = (0::rat)"  by (simp add: inverse_rat_def)
   1.636 +  show "inverse 0 = (0::rat)"
   1.637 +    by (simp add: zero_rat Fract_def inverse_rat_def
   1.638 +                  inverse_congruent UN_ratrel)
   1.639  qed
   1.640  
   1.641  
   1.642  subsection {* Various Other Results *}
   1.643  
   1.644  lemma minus_rat_cancel [simp]: "b \<noteq> 0 ==> Fract (-a) (-b) = Fract a b"
   1.645 -by (simp add: Fract_equality eq_fraction_iff)
   1.646 +by (simp add: eq_rat)
   1.647  
   1.648  theorem Rat_induct_pos [case_names Fract, induct type: rat]:
   1.649    assumes step: "!!a b. 0 < b ==> P (Fract a b)"