Up to index of Isabelle/HOL/HOL-Complex
theory Rational(* Title: HOL/Library/Rational.thy
ID: $Id: Rational.thy,v 1.11 2005/06/17 14:13:09 haftmann Exp $
Author: Markus Wenzel, TU Muenchen
*)
header {* Rational numbers *}
theory Rational
imports Quotient
uses ("rat_arith.ML")
begin
subsection {* Fractions *}
subsubsection {* The type of fractions *}
typedef fraction = "{(a, b) :: int × int | a b. b ≠ 0}"
proof
show "(0, 1) ∈ ?fraction" by simp
qed
constdefs
fract :: "int => int => fraction"
"fract a b == Abs_fraction (a, b)"
num :: "fraction => int"
"num Q == fst (Rep_fraction Q)"
den :: "fraction => int"
"den Q == snd (Rep_fraction Q)"
lemma fract_num [simp]: "b ≠ 0 ==> num (fract a b) = a"
by (simp add: fract_def num_def fraction_def Abs_fraction_inverse)
lemma fract_den [simp]: "b ≠ 0 ==> den (fract a b) = b"
by (simp add: fract_def den_def fraction_def Abs_fraction_inverse)
lemma fraction_cases [case_names fract, cases type: fraction]:
"(!!a b. Q = fract a b ==> b ≠ 0 ==> C) ==> C"
proof -
assume r: "!!a b. Q = fract a b ==> b ≠ 0 ==> C"
obtain a b where "Q = fract a b" and "b ≠ 0"
by (cases Q) (auto simp add: fract_def fraction_def)
thus C by (rule r)
qed
lemma fraction_induct [case_names fract, induct type: fraction]:
"(!!a b. b ≠ 0 ==> P (fract a b)) ==> P Q"
by (cases Q) simp
subsubsection {* Equivalence of fractions *}
instance fraction :: eqv ..
defs (overloaded)
equiv_fraction_def: "Q ∼ R == num Q * den R = num R * den Q"
lemma equiv_fraction_iff [iff]:
"b ≠ 0 ==> b' ≠ 0 ==> (fract a b ∼ fract a' b') = (a * b' = a' * b)"
by (simp add: equiv_fraction_def)
instance fraction :: equiv
proof
fix Q R S :: fraction
{
show "Q ∼ Q"
proof (induct Q)
fix a b :: int
assume "b ≠ 0" and "b ≠ 0"
with refl show "fract a b ∼ fract a b" ..
qed
next
assume "Q ∼ R" and "R ∼ S"
show "Q ∼ S"
proof (insert prems, induct Q, induct R, induct S)
fix a b a' b' a'' b'' :: int
assume b: "b ≠ 0" and b': "b' ≠ 0" and b'': "b'' ≠ 0"
assume "fract a b ∼ fract a' b'" hence eq1: "a * b' = a' * b" ..
assume "fract a' b' ∼ fract a'' b''" hence eq2: "a' * b'' = a'' * b'" ..
have "a * b'' = a'' * b"
proof cases
assume "a' = 0"
with b' eq1 eq2 have "a = 0 ∧ a'' = 0" by auto
thus ?thesis by simp
next
assume a': "a' ≠ 0"
from eq1 eq2 have "(a * b') * (a' * b'') = (a' * b) * (a'' * b')" by simp
hence "(a * b'') * (a' * b') = (a'' * b) * (a' * b')" by (simp only: mult_ac)
with a' b' show ?thesis by simp
qed
thus "fract a b ∼ fract a'' b''" ..
qed
next
show "Q ∼ R ==> R ∼ Q"
proof (induct Q, induct R)
fix a b a' b' :: int
assume b: "b ≠ 0" and b': "b' ≠ 0"
assume "fract a b ∼ fract a' b'"
hence "a * b' = a' * b" ..
hence "a' * b = a * b'" ..
thus "fract a' b' ∼ fract a b" ..
qed
}
qed
lemma eq_fraction_iff [iff]:
"b ≠ 0 ==> b' ≠ 0 ==> (⌊fract a b⌋ = ⌊fract a' b'⌋) = (a * b' = a' * b)"
by (simp add: equiv_fraction_iff quot_equality)
subsubsection {* Operations on fractions *}
text {*
We define the basic arithmetic operations on fractions and
demonstrate their ``well-definedness'', i.e.\ congruence with respect
to equivalence of fractions.
*}
instance fraction :: "{zero, one, plus, minus, times, inverse, ord}" ..
defs (overloaded)
zero_fraction_def: "0 == fract 0 1"
one_fraction_def: "1 == fract 1 1"
add_fraction_def: "Q + R ==
fract (num Q * den R + num R * den Q) (den Q * den R)"
minus_fraction_def: "-Q == fract (-(num Q)) (den Q)"
mult_fraction_def: "Q * R == fract (num Q * num R) (den Q * den R)"
inverse_fraction_def: "inverse Q == fract (den Q) (num Q)"
le_fraction_def: "Q ≤ R ==
(num Q * den R) * (den Q * den R) ≤ (num R * den Q) * (den Q * den R)"
lemma is_zero_fraction_iff: "b ≠ 0 ==> (⌊fract a b⌋ = ⌊0⌋) = (a = 0)"
by (simp add: zero_fraction_def eq_fraction_iff)
theorem add_fraction_cong:
"⌊fract a b⌋ = ⌊fract a' b'⌋ ==> ⌊fract c d⌋ = ⌊fract c' d'⌋
==> b ≠ 0 ==> b' ≠ 0 ==> d ≠ 0 ==> d' ≠ 0
==> ⌊fract a b + fract c d⌋ = ⌊fract a' b' + fract c' d'⌋"
proof -
assume neq: "b ≠ 0" "b' ≠ 0" "d ≠ 0" "d' ≠ 0"
assume "⌊fract a b⌋ = ⌊fract a' b'⌋" hence eq1: "a * b' = a' * b" ..
assume "⌊fract c d⌋ = ⌊fract c' d'⌋" hence eq2: "c * d' = c' * d" ..
have "⌊fract (a * d + c * b) (b * d)⌋ = ⌊fract (a' * d' + c' * b') (b' * d')⌋"
proof
show "(a * d + c * b) * (b' * d') = (a' * d' + c' * b') * (b * d)"
(is "?lhs = ?rhs")
proof -
have "?lhs = (a * b') * (d * d') + (c * d') * (b * b')"
by (simp add: int_distrib mult_ac)
also have "... = (a' * b) * (d * d') + (c' * d) * (b * b')"
by (simp only: eq1 eq2)
also have "... = ?rhs"
by (simp add: int_distrib mult_ac)
finally show "?lhs = ?rhs" .
qed
from neq show "b * d ≠ 0" by simp
from neq show "b' * d' ≠ 0" by simp
qed
with neq show ?thesis by (simp add: add_fraction_def)
qed
theorem minus_fraction_cong:
"⌊fract a b⌋ = ⌊fract a' b'⌋ ==> b ≠ 0 ==> b' ≠ 0
==> ⌊-(fract a b)⌋ = ⌊-(fract a' b')⌋"
proof -
assume neq: "b ≠ 0" "b' ≠ 0"
assume "⌊fract a b⌋ = ⌊fract a' b'⌋"
hence "a * b' = a' * b" ..
hence "-a * b' = -a' * b" by simp
hence "⌊fract (-a) b⌋ = ⌊fract (-a') b'⌋" ..
with neq show ?thesis by (simp add: minus_fraction_def)
qed
theorem mult_fraction_cong:
"⌊fract a b⌋ = ⌊fract a' b'⌋ ==> ⌊fract c d⌋ = ⌊fract c' d'⌋
==> b ≠ 0 ==> b' ≠ 0 ==> d ≠ 0 ==> d' ≠ 0
==> ⌊fract a b * fract c d⌋ = ⌊fract a' b' * fract c' d'⌋"
proof -
assume neq: "b ≠ 0" "b' ≠ 0" "d ≠ 0" "d' ≠ 0"
assume "⌊fract a b⌋ = ⌊fract a' b'⌋" hence eq1: "a * b' = a' * b" ..
assume "⌊fract c d⌋ = ⌊fract c' d'⌋" hence eq2: "c * d' = c' * d" ..
have "⌊fract (a * c) (b * d)⌋ = ⌊fract (a' * c') (b' * d')⌋"
proof
from eq1 eq2 have "(a * b') * (c * d') = (a' * b) * (c' * d)" by simp
thus "(a * c) * (b' * d') = (a' * c') * (b * d)" by (simp add: mult_ac)
from neq show "b * d ≠ 0" by simp
from neq show "b' * d' ≠ 0" by simp
qed
with neq show "⌊fract a b * fract c d⌋ = ⌊fract a' b' * fract c' d'⌋"
by (simp add: mult_fraction_def)
qed
theorem inverse_fraction_cong:
"⌊fract a b⌋ = ⌊fract a' b'⌋ ==> ⌊fract a b⌋ ≠ ⌊0⌋ ==> ⌊fract a' b'⌋ ≠ ⌊0⌋
==> b ≠ 0 ==> b' ≠ 0
==> ⌊inverse (fract a b)⌋ = ⌊inverse (fract a' b')⌋"
proof -
assume neq: "b ≠ 0" "b' ≠ 0"
assume "⌊fract a b⌋ ≠ ⌊0⌋" and "⌊fract a' b'⌋ ≠ ⌊0⌋"
with neq obtain "a ≠ 0" and "a' ≠ 0" by (simp add: is_zero_fraction_iff)
assume "⌊fract a b⌋ = ⌊fract a' b'⌋"
hence "a * b' = a' * b" ..
hence "b * a' = b' * a" by (simp only: mult_ac)
hence "⌊fract b a⌋ = ⌊fract b' a'⌋" ..
with neq show ?thesis by (simp add: inverse_fraction_def)
qed
theorem le_fraction_cong:
"⌊fract a b⌋ = ⌊fract a' b'⌋ ==> ⌊fract c d⌋ = ⌊fract c' d'⌋
==> b ≠ 0 ==> b' ≠ 0 ==> d ≠ 0 ==> d' ≠ 0
==> (fract a b ≤ fract c d) = (fract a' b' ≤ fract c' d')"
proof -
assume neq: "b ≠ 0" "b' ≠ 0" "d ≠ 0" "d' ≠ 0"
assume "⌊fract a b⌋ = ⌊fract a' b'⌋" hence eq1: "a * b' = a' * b" ..
assume "⌊fract c d⌋ = ⌊fract c' d'⌋" hence eq2: "c * d' = c' * d" ..
let ?le = "λa b c d. ((a * d) * (b * d) ≤ (c * b) * (b * d))"
{
fix a b c d x :: int assume x: "x ≠ 0"
have "?le a b c d = ?le (a * x) (b * x) c d"
proof -
from x have "0 < x * x" by (auto simp add: zero_less_mult_iff)
hence "?le a b c d =
((a * d) * (b * d) * (x * x) ≤ (c * b) * (b * d) * (x * x))"
by (simp add: mult_le_cancel_right)
also have "... = ?le (a * x) (b * x) c d"
by (simp add: mult_ac)
finally show ?thesis .
qed
} note le_factor = this
let ?D = "b * d" and ?D' = "b' * d'"
from neq have D: "?D ≠ 0" by simp
from neq have "?D' ≠ 0" by simp
hence "?le a b c d = ?le (a * ?D') (b * ?D') c d"
by (rule le_factor)
also have "... = ((a * b') * ?D * ?D' * d * d' ≤ (c * d') * ?D * ?D' * b * b')"
by (simp add: mult_ac)
also have "... = ((a' * b) * ?D * ?D' * d * d' ≤ (c' * d) * ?D * ?D' * b * b')"
by (simp only: eq1 eq2)
also have "... = ?le (a' * ?D) (b' * ?D) c' d'"
by (simp add: mult_ac)
also from D have "... = ?le a' b' c' d'"
by (rule le_factor [symmetric])
finally have "?le a b c d = ?le a' b' c' d'" .
with neq show ?thesis by (simp add: le_fraction_def)
qed
subsection {* Rational numbers *}
subsubsection {* The type of rational numbers *}
typedef (Rat)
rat = "UNIV :: fraction quot set" ..
lemma RatI [intro, simp]: "Q ∈ Rat"
by (simp add: Rat_def)
constdefs
fraction_of :: "rat => fraction"
"fraction_of q == pick (Rep_Rat q)"
rat_of :: "fraction => rat"
"rat_of Q == Abs_Rat ⌊Q⌋"
theorem rat_of_equality [iff?]: "(rat_of Q = rat_of Q') = (⌊Q⌋ = ⌊Q'⌋)"
by (simp add: rat_of_def Abs_Rat_inject)
lemma rat_of: "⌊Q⌋ = ⌊Q'⌋ ==> rat_of Q = rat_of Q'" ..
constdefs
Fract :: "int => int => rat"
"Fract a b == rat_of (fract a b)"
theorem Fract_inverse: "⌊fraction_of (Fract a b)⌋ = ⌊fract a b⌋"
by (simp add: fraction_of_def rat_of_def Fract_def Abs_Rat_inverse pick_inverse)
theorem Fract_equality [iff?]:
"(Fract a b = Fract c d) = (⌊fract a b⌋ = ⌊fract c d⌋)"
by (simp add: Fract_def rat_of_equality)
theorem eq_rat:
"b ≠ 0 ==> d ≠ 0 ==> (Fract a b = Fract c d) = (a * d = c * b)"
by (simp add: Fract_equality eq_fraction_iff)
theorem Rat_cases [case_names Fract, cases type: rat]:
"(!!a b. q = Fract a b ==> b ≠ 0 ==> C) ==> C"
proof -
assume r: "!!a b. q = Fract a b ==> b ≠ 0 ==> C"
obtain x where "q = Abs_Rat x" by (cases q)
moreover obtain Q where "x = ⌊Q⌋" by (cases x)
moreover obtain a b where "Q = fract a b" and "b ≠ 0" by (cases Q)
ultimately have "q = Fract a b" by (simp only: Fract_def rat_of_def)
thus ?thesis by (rule r)
qed
theorem Rat_induct [case_names Fract, induct type: rat]:
"(!!a b. b ≠ 0 ==> P (Fract a b)) ==> P q"
by (cases q) simp
subsubsection {* Canonical function definitions *}
text {*
Note that the unconditional version below is much easier to read.
*}
theorem rat_cond_function:
"(!!q r. P ⌊fraction_of q⌋ ⌊fraction_of r⌋ ==>
f q r == g (fraction_of q) (fraction_of r)) ==>
(!!a b a' b' c d c' d'.
⌊fract a b⌋ = ⌊fract a' b'⌋ ==> ⌊fract c d⌋ = ⌊fract c' d'⌋ ==>
P ⌊fract a b⌋ ⌊fract c d⌋ ==> P ⌊fract a' b'⌋ ⌊fract c' d'⌋ ==>
b ≠ 0 ==> b' ≠ 0 ==> d ≠ 0 ==> d' ≠ 0 ==>
g (fract a b) (fract c d) = g (fract a' b') (fract c' d')) ==>
P ⌊fract a b⌋ ⌊fract c d⌋ ==>
f (Fract a b) (Fract c d) = g (fract a b) (fract c d)"
(is "PROP ?eq ==> PROP ?cong ==> ?P ==> _")
proof -
assume eq: "PROP ?eq" and cong: "PROP ?cong" and P: ?P
have "f (Abs_Rat ⌊fract a b⌋) (Abs_Rat ⌊fract c d⌋) = g (fract a b) (fract c d)"
proof (rule quot_cond_function)
fix X Y assume "P X Y"
with eq show "f (Abs_Rat X) (Abs_Rat Y) == g (pick X) (pick Y)"
by (simp add: fraction_of_def pick_inverse Abs_Rat_inverse)
next
fix Q Q' R R' :: fraction
show "⌊Q⌋ = ⌊Q'⌋ ==> ⌊R⌋ = ⌊R'⌋ ==>
P ⌊Q⌋ ⌊R⌋ ==> P ⌊Q'⌋ ⌊R'⌋ ==> g Q R = g Q' R'"
by (induct Q, induct Q', induct R, induct R') (rule cong)
qed
thus ?thesis by (unfold Fract_def rat_of_def)
qed
theorem rat_function:
"(!!q r. f q r == g (fraction_of q) (fraction_of r)) ==>
(!!a b a' b' c d c' d'.
⌊fract a b⌋ = ⌊fract a' b'⌋ ==> ⌊fract c d⌋ = ⌊fract c' d'⌋ ==>
b ≠ 0 ==> b' ≠ 0 ==> d ≠ 0 ==> d' ≠ 0 ==>
g (fract a b) (fract c d) = g (fract a' b') (fract c' d')) ==>
f (Fract a b) (Fract c d) = g (fract a b) (fract c d)"
proof -
case rule_context from this TrueI
show ?thesis by (rule rat_cond_function)
qed
subsubsection {* Standard operations on rational numbers *}
instance rat :: "{zero, one, plus, minus, times, inverse, ord}" ..
defs (overloaded)
zero_rat_def: "0 == rat_of 0"
one_rat_def: "1 == rat_of 1"
add_rat_def: "q + r == rat_of (fraction_of q + fraction_of r)"
minus_rat_def: "-q == rat_of (-(fraction_of q))"
diff_rat_def: "q - r == q + (-(r::rat))"
mult_rat_def: "q * r == rat_of (fraction_of q * fraction_of r)"
inverse_rat_def: "inverse q ==
if q=0 then 0 else rat_of (inverse (fraction_of q))"
divide_rat_def: "q / r == q * inverse (r::rat)"
le_rat_def: "q ≤ r == fraction_of q ≤ fraction_of r"
less_rat_def: "q < r == q ≤ r ∧ q ≠ (r::rat)"
abs_rat_def: "¦q¦ == if q < 0 then -q else (q::rat)"
theorem zero_rat: "0 = Fract 0 1"
by (simp add: zero_rat_def zero_fraction_def rat_of_def Fract_def)
theorem one_rat: "1 = Fract 1 1"
by (simp add: one_rat_def one_fraction_def rat_of_def Fract_def)
theorem add_rat: "b ≠ 0 ==> d ≠ 0 ==>
Fract a b + Fract c d = Fract (a * d + c * b) (b * d)"
proof -
have "Fract a b + Fract c d = rat_of (fract a b + fract c d)"
by (rule rat_function, rule add_rat_def, rule rat_of, rule add_fraction_cong)
also
assume "b ≠ 0" "d ≠ 0"
hence "fract a b + fract c d = fract (a * d + c * b) (b * d)"
by (simp add: add_fraction_def)
finally show ?thesis by (unfold Fract_def)
qed
theorem minus_rat: "b ≠ 0 ==> -(Fract a b) = Fract (-a) b"
proof -
have "-(Fract a b) = rat_of (-(fract a b))"
by (rule rat_function, rule minus_rat_def, rule rat_of, rule minus_fraction_cong)
also assume "b ≠ 0" hence "-(fract a b) = fract (-a) b"
by (simp add: minus_fraction_def)
finally show ?thesis by (unfold Fract_def)
qed
theorem diff_rat: "b ≠ 0 ==> d ≠ 0 ==>
Fract a b - Fract c d = Fract (a * d - c * b) (b * d)"
by (simp add: diff_rat_def add_rat minus_rat)
theorem mult_rat: "b ≠ 0 ==> d ≠ 0 ==>
Fract a b * Fract c d = Fract (a * c) (b * d)"
proof -
have "Fract a b * Fract c d = rat_of (fract a b * fract c d)"
by (rule rat_function, rule mult_rat_def, rule rat_of, rule mult_fraction_cong)
also
assume "b ≠ 0" "d ≠ 0"
hence "fract a b * fract c d = fract (a * c) (b * d)"
by (simp add: mult_fraction_def)
finally show ?thesis by (unfold Fract_def)
qed
theorem inverse_rat: "Fract a b ≠ 0 ==> b ≠ 0 ==>
inverse (Fract a b) = Fract b a"
proof -
assume neq: "b ≠ 0" and nonzero: "Fract a b ≠ 0"
hence "⌊fract a b⌋ ≠ ⌊0⌋"
by (simp add: zero_rat eq_rat is_zero_fraction_iff)
with _ inverse_fraction_cong [THEN rat_of]
have "inverse (Fract a b) = rat_of (inverse (fract a b))"
proof (rule rat_cond_function)
fix q assume cond: "⌊fraction_of q⌋ ≠ ⌊0⌋"
have "q ≠ 0"
proof (cases q)
fix a b assume "b ≠ 0" and "q = Fract a b"
from this cond show ?thesis
by (simp add: Fract_inverse is_zero_fraction_iff zero_rat eq_rat)
qed
thus "inverse q == rat_of (inverse (fraction_of q))"
by (simp add: inverse_rat_def)
qed
also from neq nonzero have "inverse (fract a b) = fract b a"
by (simp add: inverse_fraction_def)
finally show ?thesis by (unfold Fract_def)
qed
theorem divide_rat: "Fract c d ≠ 0 ==> b ≠ 0 ==> d ≠ 0 ==>
Fract a b / Fract c d = Fract (a * d) (b * c)"
proof -
assume neq: "b ≠ 0" "d ≠ 0" and nonzero: "Fract c d ≠ 0"
hence "c ≠ 0" by (simp add: zero_rat eq_rat)
with neq nonzero show ?thesis
by (simp add: divide_rat_def inverse_rat mult_rat)
qed
theorem le_rat: "b ≠ 0 ==> d ≠ 0 ==>
(Fract a b ≤ Fract c d) = ((a * d) * (b * d) ≤ (c * b) * (b * d))"
proof -
have "(Fract a b ≤ Fract c d) = (fract a b ≤ fract c d)"
by (rule rat_function, rule le_rat_def, rule le_fraction_cong)
also
assume "b ≠ 0" "d ≠ 0"
hence "(fract a b ≤ fract c d) = ((a * d) * (b * d) ≤ (c * b) * (b * d))"
by (simp add: le_fraction_def)
finally show ?thesis .
qed
theorem less_rat: "b ≠ 0 ==> d ≠ 0 ==>
(Fract a b < Fract c d) = ((a * d) * (b * d) < (c * b) * (b * d))"
by (simp add: less_rat_def le_rat eq_rat order_less_le)
theorem abs_rat: "b ≠ 0 ==> ¦Fract a b¦ = Fract ¦a¦ ¦b¦"
by (simp add: abs_rat_def minus_rat zero_rat less_rat eq_rat)
(auto simp add: mult_less_0_iff zero_less_mult_iff order_le_less
split: abs_split)
subsubsection {* The ordered field of rational numbers *}
lemma rat_add_assoc: "(q + r) + s = q + (r + (s::rat))"
by (induct q, induct r, induct s)
(simp add: add_rat add_ac mult_ac int_distrib)
lemma rat_add_0: "0 + q = (q::rat)"
by (induct q) (simp add: zero_rat add_rat)
lemma rat_left_minus: "(-q) + q = (0::rat)"
by (induct q) (simp add: zero_rat minus_rat add_rat eq_rat)
instance rat :: field
proof
fix q r s :: rat
show "(q + r) + s = q + (r + s)"
by (rule rat_add_assoc)
show "q + r = r + q"
by (induct q, induct r) (simp add: add_rat add_ac mult_ac)
show "0 + q = q"
by (induct q) (simp add: zero_rat add_rat)
show "(-q) + q = 0"
by (rule rat_left_minus)
show "q - r = q + (-r)"
by (induct q, induct r) (simp add: add_rat minus_rat diff_rat)
show "(q * r) * s = q * (r * s)"
by (induct q, induct r, induct s) (simp add: mult_rat mult_ac)
show "q * r = r * q"
by (induct q, induct r) (simp add: mult_rat mult_ac)
show "1 * q = q"
by (induct q) (simp add: one_rat mult_rat)
show "(q + r) * s = q * s + r * s"
by (induct q, induct r, induct s)
(simp add: add_rat mult_rat eq_rat int_distrib)
show "q ≠ 0 ==> inverse q * q = 1"
by (induct q) (simp add: inverse_rat mult_rat one_rat zero_rat eq_rat)
show "q / r = q * inverse r"
by (simp add: divide_rat_def)
show "0 ≠ (1::rat)"
by (simp add: zero_rat one_rat eq_rat)
qed
instance rat :: linorder
proof
fix q r s :: rat
{
assume "q ≤ r" and "r ≤ s"
show "q ≤ s"
proof (insert prems, induct q, induct r, induct s)
fix a b c d e f :: int
assume neq: "b ≠ 0" "d ≠ 0" "f ≠ 0"
assume 1: "Fract a b ≤ Fract c d" and 2: "Fract c d ≤ Fract e f"
show "Fract a b ≤ Fract e f"
proof -
from neq obtain bb: "0 < b * b" and dd: "0 < d * d" and ff: "0 < f * f"
by (auto simp add: zero_less_mult_iff linorder_neq_iff)
have "(a * d) * (b * d) * (f * f) ≤ (c * b) * (b * d) * (f * f)"
proof -
from neq 1 have "(a * d) * (b * d) ≤ (c * b) * (b * d)"
by (simp add: le_rat)
with ff show ?thesis by (simp add: mult_le_cancel_right)
qed
also have "... = (c * f) * (d * f) * (b * b)"
by (simp only: mult_ac)
also have "... ≤ (e * d) * (d * f) * (b * b)"
proof -
from neq 2 have "(c * f) * (d * f) ≤ (e * d) * (d * f)"
by (simp add: le_rat)
with bb show ?thesis by (simp add: mult_le_cancel_right)
qed
finally have "(a * f) * (b * f) * (d * d) ≤ e * b * (b * f) * (d * d)"
by (simp only: mult_ac)
with dd have "(a * f) * (b * f) ≤ (e * b) * (b * f)"
by (simp add: mult_le_cancel_right)
with neq show ?thesis by (simp add: le_rat)
qed
qed
next
assume "q ≤ r" and "r ≤ q"
show "q = r"
proof (insert prems, induct q, induct r)
fix a b c d :: int
assume neq: "b ≠ 0" "d ≠ 0"
assume 1: "Fract a b ≤ Fract c d" and 2: "Fract c d ≤ Fract a b"
show "Fract a b = Fract c d"
proof -
from neq 1 have "(a * d) * (b * d) ≤ (c * b) * (b * d)"
by (simp add: le_rat)
also have "... ≤ (a * d) * (b * d)"
proof -
from neq 2 have "(c * b) * (d * b) ≤ (a * d) * (d * b)"
by (simp add: le_rat)
thus ?thesis by (simp only: mult_ac)
qed
finally have "(a * d) * (b * d) = (c * b) * (b * d)" .
moreover from neq have "b * d ≠ 0" by simp
ultimately have "a * d = c * b" by simp
with neq show ?thesis by (simp add: eq_rat)
qed
qed
next
show "q ≤ q"
by (induct q) (simp add: le_rat)
show "(q < r) = (q ≤ r ∧ q ≠ r)"
by (simp only: less_rat_def)
show "q ≤ r ∨ r ≤ q"
by (induct q, induct r) (simp add: le_rat mult_ac, arith)
}
qed
instance rat :: ordered_field
proof
fix q r s :: rat
show "q ≤ r ==> s + q ≤ s + r"
proof (induct q, induct r, induct s)
fix a b c d e f :: int
assume neq: "b ≠ 0" "d ≠ 0" "f ≠ 0"
assume le: "Fract a b ≤ Fract c d"
show "Fract e f + Fract a b ≤ Fract e f + Fract c d"
proof -
let ?F = "f * f" from neq have F: "0 < ?F"
by (auto simp add: zero_less_mult_iff)
from neq le have "(a * d) * (b * d) ≤ (c * b) * (b * d)"
by (simp add: le_rat)
with F have "(a * d) * (b * d) * ?F * ?F ≤ (c * b) * (b * d) * ?F * ?F"
by (simp add: mult_le_cancel_right)
with neq show ?thesis by (simp add: add_rat le_rat mult_ac int_distrib)
qed
qed
show "q < r ==> 0 < s ==> s * q < s * r"
proof (induct q, induct r, induct s)
fix a b c d e f :: int
assume neq: "b ≠ 0" "d ≠ 0" "f ≠ 0"
assume le: "Fract a b < Fract c d"
assume gt: "0 < Fract e f"
show "Fract e f * Fract a b < Fract e f * Fract c d"
proof -
let ?E = "e * f" and ?F = "f * f"
from neq gt have "0 < ?E"
by (auto simp add: zero_rat less_rat le_rat order_less_le eq_rat)
moreover from neq have "0 < ?F"
by (auto simp add: zero_less_mult_iff)
moreover from neq le have "(a * d) * (b * d) < (c * b) * (b * d)"
by (simp add: less_rat)
ultimately have "(a * d) * (b * d) * ?E * ?F < (c * b) * (b * d) * ?E * ?F"
by (simp add: mult_less_cancel_right)
with neq show ?thesis
by (simp add: less_rat mult_rat mult_ac)
qed
qed
show "¦q¦ = (if q < 0 then -q else q)"
by (simp only: abs_rat_def)
qed
instance rat :: division_by_zero
proof
show "inverse 0 = (0::rat)" by (simp add: inverse_rat_def)
qed
subsection {* Various Other Results *}
lemma minus_rat_cancel [simp]: "b ≠ 0 ==> Fract (-a) (-b) = Fract a b"
by (simp add: Fract_equality eq_fraction_iff)
theorem Rat_induct_pos [case_names Fract, induct type: rat]:
assumes step: "!!a b. 0 < b ==> P (Fract a b)"
shows "P q"
proof (cases q)
have step': "!!a b. b < 0 ==> P (Fract a b)"
proof -
fix a::int and b::int
assume b: "b < 0"
hence "0 < -b" by simp
hence "P (Fract (-a) (-b))" by (rule step)
thus "P (Fract a b)" by (simp add: order_less_imp_not_eq [OF b])
qed
case (Fract a b)
thus "P q" by (force simp add: linorder_neq_iff step step')
qed
lemma zero_less_Fract_iff:
"0 < b ==> (0 < Fract a b) = (0 < a)"
by (simp add: zero_rat less_rat order_less_imp_not_eq2 zero_less_mult_iff)
lemma Fract_add_one: "n ≠ 0 ==> Fract (m + n) n = Fract m n + 1"
apply (insert add_rat [of concl: m n 1 1])
apply (simp add: one_rat [symmetric])
done
lemma Fract_of_nat_eq: "Fract (of_nat k) 1 = of_nat k"
apply (induct k)
apply (simp add: zero_rat)
apply (simp add: Fract_add_one)
done
lemma Fract_of_int_eq: "Fract k 1 = of_int k"
proof (cases k rule: int_cases)
case (nonneg n)
thus ?thesis by (simp add: int_eq_of_nat Fract_of_nat_eq)
next
case (neg n)
hence "Fract k 1 = - (Fract (of_nat (Suc n)) 1)"
by (simp only: minus_rat int_eq_of_nat)
also have "... = - (of_nat (Suc n))"
by (simp only: Fract_of_nat_eq)
finally show ?thesis
by (simp add: only: prems int_eq_of_nat of_int_minus of_int_of_nat_eq)
qed
subsection {* Numerals and Arithmetic *}
instance rat :: number ..
defs (overloaded)
rat_number_of_def: "(number_of w :: rat) == of_int (Rep_Bin w)"
--{*the type constraint is essential!*}
instance rat :: number_ring
by (intro_classes, simp add: rat_number_of_def)
declare diff_rat_def [symmetric]
use "rat_arith.ML"
setup rat_arith_setup
end
lemma fract_num:
b ≠ 0 ==> num (fract a b) = a
lemma fract_den:
b ≠ 0 ==> den (fract a b) = b
lemma fraction_cases:
(!!a b. [| Q = fract a b; b ≠ 0 |] ==> C) ==> C
lemma fraction_induct:
(!!a b. b ≠ 0 ==> P (fract a b)) ==> P Q
lemma equiv_fraction_iff:
[| b ≠ 0; b' ≠ 0 |] ==> (fract a b ∼ fract a' b') = (a * b' = a' * b)
lemma eq_fraction_iff:
[| b ≠ 0; b' ≠ 0 |] ==> (⌊fract a b⌋ = ⌊fract a' b'⌋) = (a * b' = a' * b)
lemma is_zero_fraction_iff:
b ≠ 0 ==> (⌊fract a b⌋ = ⌊0⌋) = (a = 0)
theorem add_fraction_cong:
[| ⌊fract a b⌋ = ⌊fract a' b'⌋; ⌊fract c d⌋ = ⌊fract c' d'⌋; b ≠ 0; b' ≠ 0; d ≠ 0; d' ≠ 0 |] ==> ⌊fract a b + fract c d⌋ = ⌊fract a' b' + fract c' d'⌋
theorem minus_fraction_cong:
[| ⌊fract a b⌋ = ⌊fract a' b'⌋; b ≠ 0; b' ≠ 0 |] ==> ⌊- fract a b⌋ = ⌊- fract a' b'⌋
theorem mult_fraction_cong:
[| ⌊fract a b⌋ = ⌊fract a' b'⌋; ⌊fract c d⌋ = ⌊fract c' d'⌋; b ≠ 0; b' ≠ 0; d ≠ 0; d' ≠ 0 |] ==> ⌊fract a b * fract c d⌋ = ⌊fract a' b' * fract c' d'⌋
theorem inverse_fraction_cong:
[| ⌊fract a b⌋ = ⌊fract a' b'⌋; ⌊fract a b⌋ ≠ ⌊0⌋; ⌊fract a' b'⌋ ≠ ⌊0⌋; b ≠ 0; b' ≠ 0 |] ==> ⌊inverse (fract a b)⌋ = ⌊inverse (fract a' b')⌋
theorem le_fraction_cong:
[| ⌊fract a b⌋ = ⌊fract a' b'⌋; ⌊fract c d⌋ = ⌊fract c' d'⌋; b ≠ 0; b' ≠ 0; d ≠ 0; d' ≠ 0 |] ==> (fract a b ≤ fract c d) = (fract a' b' ≤ fract c' d')
lemma RatI:
Q ∈ Rat
theorem rat_of_equality:
(rat_of Q = rat_of Q') = (⌊Q⌋ = ⌊Q'⌋)
lemma rat_of:
⌊Q⌋ = ⌊Q'⌋ ==> rat_of Q = rat_of Q'
theorem Fract_inverse:
⌊fraction_of (Fract a b)⌋ = ⌊fract a b⌋
theorem Fract_equality:
(Fract a b = Fract c d) = (⌊fract a b⌋ = ⌊fract c d⌋)
theorem eq_rat:
[| b ≠ 0; d ≠ 0 |] ==> (Fract a b = Fract c d) = (a * d = c * b)
theorem Rat_cases:
(!!a b. [| q = Fract a b; b ≠ 0 |] ==> C) ==> C
theorem Rat_induct:
(!!a b. b ≠ 0 ==> P (Fract a b)) ==> P q
theorem rat_cond_function:
[| !!q r. P ⌊fraction_of q⌋ ⌊fraction_of r⌋ ==> f q r == g (fraction_of q) (fraction_of r); !!a b a' b' c d c' d'. [| ⌊fract a b⌋ = ⌊fract a' b'⌋; ⌊fract c d⌋ = ⌊fract c' d'⌋; P ⌊fract a b⌋ ⌊fract c d⌋; P ⌊fract a' b'⌋ ⌊fract c' d'⌋; b ≠ 0; b' ≠ 0; d ≠ 0; d' ≠ 0 |] ==> g (fract a b) (fract c d) = g (fract a' b') (fract c' d'); P ⌊fract a b⌋ ⌊fract c d⌋ |] ==> f (Fract a b) (Fract c d) = g (fract a b) (fract c d)
theorem rat_function:
[| !!q r. f q r == g (fraction_of q) (fraction_of r); !!a b a' b' c d c' d'. [| ⌊fract a b⌋ = ⌊fract a' b'⌋; ⌊fract c d⌋ = ⌊fract c' d'⌋; b ≠ 0; b' ≠ 0; d ≠ 0; d' ≠ 0 |] ==> g (fract a b) (fract c d) = g (fract a' b') (fract c' d') |] ==> f (Fract a b) (Fract c d) = g (fract a b) (fract c d)
theorem zero_rat:
0 = Fract 0 1
theorem one_rat:
1 = Fract 1 1
theorem add_rat:
[| b ≠ 0; d ≠ 0 |] ==> Fract a b + Fract c d = Fract (a * d + c * b) (b * d)
theorem minus_rat:
b ≠ 0 ==> - Fract a b = Fract (- a) b
theorem diff_rat:
[| b ≠ 0; d ≠ 0 |] ==> Fract a b - Fract c d = Fract (a * d - c * b) (b * d)
theorem mult_rat:
[| b ≠ 0; d ≠ 0 |] ==> Fract a b * Fract c d = Fract (a * c) (b * d)
theorem inverse_rat:
[| Fract a b ≠ 0; b ≠ 0 |] ==> inverse (Fract a b) = Fract b a
theorem divide_rat:
[| Fract c d ≠ 0; b ≠ 0; d ≠ 0 |] ==> Fract a b / Fract c d = Fract (a * d) (b * c)
theorem le_rat:
[| b ≠ 0; d ≠ 0 |] ==> (Fract a b ≤ Fract c d) = (a * d * (b * d) ≤ c * b * (b * d))
theorem less_rat:
[| b ≠ 0; d ≠ 0 |] ==> (Fract a b < Fract c d) = (a * d * (b * d) < c * b * (b * d))
theorem abs_rat:
b ≠ 0 ==> ¦Fract a b¦ = Fract ¦a¦ ¦b¦
lemma rat_add_assoc:
q + r + s = q + (r + s)
lemma rat_add_0:
0 + q = q
lemma rat_left_minus:
- q + q = 0
lemma minus_rat_cancel:
b ≠ 0 ==> Fract (- a) (- b) = Fract a b
theorem Rat_induct_pos:
(!!a b. 0 < b ==> P (Fract a b)) ==> P q
lemma zero_less_Fract_iff:
0 < b ==> (0 < Fract a b) = (0 < a)
lemma Fract_add_one:
n ≠ 0 ==> Fract (m + n) n = Fract m n + 1
lemma Fract_of_nat_eq:
Fract (of_nat k) 1 = of_nat k
lemma Fract_of_int_eq:
Fract k 1 = of_int k