| author | Andreas Lochbihler | 
| Tue, 07 Jun 2016 15:12:27 +0200 | |
| changeset 63243 | 1bc6816fd525 | 
| parent 63092 | a949b2a5f51d | 
| child 64290 | fb5c74a58796 | 
| permissions | -rw-r--r-- | 
| 35372 | 1  | 
(* Title: HOL/Library/Fraction_Field.thy  | 
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
2  | 
Author: Amine Chaieb, University of Cambridge  | 
| 
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
3  | 
*)  | 
| 
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
4  | 
|
| 60500 | 5  | 
section\<open>A formalization of the fraction field of any integral domain;  | 
6  | 
generalization of theory Rat from int to any integral domain\<close>  | 
|
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
7  | 
|
| 
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
8  | 
theory Fraction_Field  | 
| 35372 | 9  | 
imports Main  | 
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
10  | 
begin  | 
| 
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
11  | 
|
| 60500 | 12  | 
subsection \<open>General fractions construction\<close>  | 
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
13  | 
|
| 60500 | 14  | 
subsubsection \<open>Construction of the type of fractions\<close>  | 
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
15  | 
|
| 61106 | 16  | 
context idom begin  | 
17  | 
||
18  | 
definition fractrel :: "'a \<times> 'a \<Rightarrow> 'a * 'a \<Rightarrow> bool" where  | 
|
19  | 
"fractrel = (\<lambda>x y. snd x \<noteq> 0 \<and> snd y \<noteq> 0 \<and> fst x * snd y = fst y * snd x)"  | 
|
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
20  | 
|
| 
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
21  | 
lemma fractrel_iff [simp]:  | 
| 61106 | 22  | 
"fractrel x y \<longleftrightarrow> snd x \<noteq> 0 \<and> snd y \<noteq> 0 \<and> fst x * snd y = fst y * snd x"  | 
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
23  | 
by (simp add: fractrel_def)  | 
| 
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
24  | 
|
| 61106 | 25  | 
lemma symp_fractrel: "symp fractrel"  | 
26  | 
by (simp add: symp_def)  | 
|
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
27  | 
|
| 61106 | 28  | 
lemma transp_fractrel: "transp fractrel"  | 
29  | 
proof (rule transpI, unfold split_paired_all)  | 
|
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
30  | 
fix a b a' b' a'' b'' :: 'a  | 
| 61106 | 31  | 
assume A: "fractrel (a, b) (a', b')"  | 
32  | 
assume B: "fractrel (a', b') (a'', b'')"  | 
|
| 
57514
 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 
haftmann 
parents: 
57512 
diff
changeset
 | 
33  | 
have "b' * (a * b'') = b'' * (a * b')" by (simp add: ac_simps)  | 
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
34  | 
also from A have "a * b' = a' * b" by auto  | 
| 
57514
 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 
haftmann 
parents: 
57512 
diff
changeset
 | 
35  | 
also have "b'' * (a' * b) = b * (a' * b'')" by (simp add: ac_simps)  | 
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
36  | 
also from B have "a' * b'' = a'' * b'" by auto  | 
| 
57514
 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 
haftmann 
parents: 
57512 
diff
changeset
 | 
37  | 
also have "b * (a'' * b') = b' * (a'' * b)" by (simp add: ac_simps)  | 
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
38  | 
finally have "b' * (a * b'') = b' * (a'' * b)" .  | 
| 
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
39  | 
moreover from B have "b' \<noteq> 0" by auto  | 
| 
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
40  | 
ultimately have "a * b'' = a'' * b" by simp  | 
| 61106 | 41  | 
with A B show "fractrel (a, b) (a'', b'')" by auto  | 
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
42  | 
qed  | 
| 54463 | 43  | 
|
| 61106 | 44  | 
lemma part_equivp_fractrel: "part_equivp fractrel"  | 
45  | 
using _ symp_fractrel transp_fractrel  | 
|
46  | 
by(rule part_equivpI)(rule exI[where x="(0, 1)"]; simp)  | 
|
| 
45694
 
4a8743618257
prefer typedef without extra definition and alternative name;
 
wenzelm 
parents: 
40822 
diff
changeset
 | 
47  | 
|
| 61106 | 48  | 
end  | 
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
49  | 
|
| 61260 | 50  | 
quotient_type (overloaded) 'a fract = "'a :: idom \<times> 'a" / partial: "fractrel"  | 
| 61106 | 51  | 
by(rule part_equivp_fractrel)  | 
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
52  | 
|
| 60500 | 53  | 
subsubsection \<open>Representation and basic operations\<close>  | 
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
54  | 
|
| 61106 | 55  | 
lift_definition Fract :: "'a :: idom \<Rightarrow> 'a \<Rightarrow> 'a fract"  | 
56  | 
is "\<lambda>a b. if b = 0 then (0, 1) else (a, b)"  | 
|
57  | 
by simp  | 
|
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
58  | 
|
| 53196 | 59  | 
lemma Fract_cases [cases type: fract]:  | 
60  | 
obtains (Fract) a b where "q = Fract a b" "b \<noteq> 0"  | 
|
| 61106 | 61  | 
by transfer simp  | 
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
62  | 
|
| 
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
63  | 
lemma Fract_induct [case_names Fract, induct type: fract]:  | 
| 54463 | 64  | 
"(\<And>a b. b \<noteq> 0 \<Longrightarrow> P (Fract a b)) \<Longrightarrow> P q"  | 
| 53196 | 65  | 
by (cases q) simp  | 
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
66  | 
|
| 
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
67  | 
lemma eq_fract:  | 
| 
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
68  | 
shows "\<And>a b c d. b \<noteq> 0 \<Longrightarrow> d \<noteq> 0 \<Longrightarrow> Fract a b = Fract c d \<longleftrightarrow> a * d = c * b"  | 
| 53196 | 69  | 
and "\<And>a. Fract a 0 = Fract 0 1"  | 
70  | 
and "\<And>a c. Fract 0 a = Fract 0 c"  | 
|
| 61106 | 71  | 
by(transfer; simp)+  | 
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
72  | 
|
| 53196 | 73  | 
instantiation fract :: (idom) "{comm_ring_1,power}"
 | 
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
74  | 
begin  | 
| 
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
75  | 
|
| 61106 | 76  | 
lift_definition zero_fract :: "'a fract" is "(0, 1)" by simp  | 
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
77  | 
|
| 61106 | 78  | 
lemma Zero_fract_def: "0 = Fract 0 1"  | 
79  | 
by transfer simp  | 
|
80  | 
||
81  | 
lift_definition one_fract :: "'a fract" is "(1, 1)" by simp  | 
|
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
82  | 
|
| 61106 | 83  | 
lemma One_fract_def: "1 = Fract 1 1"  | 
84  | 
by transfer simp  | 
|
85  | 
||
86  | 
lift_definition plus_fract :: "'a fract \<Rightarrow> 'a fract \<Rightarrow> 'a fract"  | 
|
87  | 
is "\<lambda>q r. (fst q * snd r + fst r * snd q, snd q * snd r)"  | 
|
88  | 
by(auto simp add: algebra_simps)  | 
|
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
89  | 
|
| 
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
90  | 
lemma add_fract [simp]:  | 
| 61106 | 91  | 
"\<lbrakk> b \<noteq> 0; d \<noteq> 0 \<rbrakk> \<Longrightarrow> Fract a b + Fract c d = Fract (a * d + c * b) (b * d)"  | 
92  | 
by transfer simp  | 
|
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
93  | 
|
| 61106 | 94  | 
lift_definition uminus_fract :: "'a fract \<Rightarrow> 'a fract"  | 
95  | 
is "\<lambda>x. (- fst x, snd x)"  | 
|
96  | 
by simp  | 
|
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
97  | 
|
| 61106 | 98  | 
lemma minus_fract [simp]:  | 
| 54463 | 99  | 
fixes a b :: "'a::idom"  | 
100  | 
shows "- Fract a b = Fract (- a) b"  | 
|
| 61106 | 101  | 
by transfer simp  | 
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
102  | 
|
| 
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
103  | 
lemma minus_fract_cancel [simp]: "Fract (- a) (- b) = Fract a b"  | 
| 
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
104  | 
by (cases "b = 0") (simp_all add: eq_fract)  | 
| 
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
105  | 
|
| 46573 | 106  | 
definition diff_fract_def: "q - r = q + - (r::'a fract)"  | 
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
107  | 
|
| 
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
108  | 
lemma diff_fract [simp]:  | 
| 61106 | 109  | 
"\<lbrakk> b \<noteq> 0; d \<noteq> 0 \<rbrakk> \<Longrightarrow> Fract a b - Fract c d = Fract (a * d - c * b) (b * d)"  | 
110  | 
by (simp add: diff_fract_def)  | 
|
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
111  | 
|
| 61106 | 112  | 
lift_definition times_fract :: "'a fract \<Rightarrow> 'a fract \<Rightarrow> 'a fract"  | 
113  | 
is "\<lambda>q r. (fst q * fst r, snd q * snd r)"  | 
|
114  | 
by(simp add: algebra_simps)  | 
|
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
115  | 
|
| 
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
116  | 
lemma mult_fract [simp]: "Fract (a::'a::idom) b * Fract c d = Fract (a * c) (b * d)"  | 
| 61106 | 117  | 
by transfer simp  | 
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
118  | 
|
| 
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
119  | 
lemma mult_fract_cancel:  | 
| 61106 | 120  | 
"c \<noteq> 0 \<Longrightarrow> Fract (c * a) (c * b) = Fract a b"  | 
121  | 
by transfer simp  | 
|
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
122  | 
|
| 47252 | 123  | 
instance  | 
124  | 
proof  | 
|
| 53196 | 125  | 
fix q r s :: "'a fract"  | 
| 54463 | 126  | 
show "(q * r) * s = q * (r * s)"  | 
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
127  | 
by (cases q, cases r, cases s) (simp add: eq_fract algebra_simps)  | 
| 53196 | 128  | 
show "q * r = r * q"  | 
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
129  | 
by (cases q, cases r) (simp add: eq_fract algebra_simps)  | 
| 53196 | 130  | 
show "1 * q = q"  | 
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
131  | 
by (cases q) (simp add: One_fract_def eq_fract)  | 
| 53196 | 132  | 
show "(q + r) + s = q + (r + s)"  | 
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
133  | 
by (cases q, cases r, cases s) (simp add: eq_fract algebra_simps)  | 
| 53196 | 134  | 
show "q + r = r + q"  | 
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
135  | 
by (cases q, cases r) (simp add: eq_fract algebra_simps)  | 
| 53196 | 136  | 
show "0 + q = q"  | 
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
137  | 
by (cases q) (simp add: Zero_fract_def eq_fract)  | 
| 53196 | 138  | 
show "- q + q = 0"  | 
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
139  | 
by (cases q) (simp add: Zero_fract_def eq_fract)  | 
| 53196 | 140  | 
show "q - r = q + - r"  | 
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
141  | 
by (cases q, cases r) (simp add: eq_fract)  | 
| 53196 | 142  | 
show "(q + r) * s = q * s + r * s"  | 
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
143  | 
by (cases q, cases r, cases s) (simp add: eq_fract algebra_simps)  | 
| 53196 | 144  | 
show "(0::'a fract) \<noteq> 1"  | 
145  | 
by (simp add: Zero_fract_def One_fract_def eq_fract)  | 
|
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
146  | 
qed  | 
| 
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
147  | 
|
| 
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
148  | 
end  | 
| 
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
149  | 
|
| 
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
150  | 
lemma of_nat_fract: "of_nat k = Fract (of_nat k) 1"  | 
| 
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
151  | 
by (induct k) (simp_all add: Zero_fract_def One_fract_def)  | 
| 
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
152  | 
|
| 
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
153  | 
lemma Fract_of_nat_eq: "Fract (of_nat k) 1 = of_nat k"  | 
| 
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
154  | 
by (rule of_nat_fract [symmetric])  | 
| 
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
155  | 
|
| 61106 | 156  | 
lemma fract_collapse:  | 
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
157  | 
"Fract 0 k = 0"  | 
| 
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
158  | 
"Fract 1 1 = 1"  | 
| 
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
159  | 
"Fract k 0 = 0"  | 
| 61106 | 160  | 
by(transfer; simp)+  | 
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
161  | 
|
| 61106 | 162  | 
lemma fract_expand:  | 
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
163  | 
"0 = Fract 0 1"  | 
| 
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
164  | 
"1 = Fract 1 1"  | 
| 
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
165  | 
by (simp_all add: fract_collapse)  | 
| 
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
166  | 
|
| 53196 | 167  | 
lemma Fract_cases_nonzero:  | 
| 54463 | 168  | 
obtains (Fract) a b where "q = Fract a b" and "b \<noteq> 0" and "a \<noteq> 0"  | 
| 53196 | 169  | 
| (0) "q = 0"  | 
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
170  | 
proof (cases "q = 0")  | 
| 53196 | 171  | 
case True  | 
172  | 
then show thesis using 0 by auto  | 
|
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
173  | 
next  | 
| 
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
174  | 
case False  | 
| 
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
175  | 
then obtain a b where "q = Fract a b" and "b \<noteq> 0" by (cases q) auto  | 
| 
53374
 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 
wenzelm 
parents: 
53196 
diff
changeset
 | 
176  | 
with False have "0 \<noteq> Fract a b" by simp  | 
| 60500 | 177  | 
with \<open>b \<noteq> 0\<close> have "a \<noteq> 0" by (simp add: Zero_fract_def eq_fract)  | 
178  | 
with Fract \<open>q = Fract a b\<close> \<open>b \<noteq> 0\<close> show thesis by auto  | 
|
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
179  | 
qed  | 
| 54463 | 180  | 
|
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
181  | 
|
| 60500 | 182  | 
subsubsection \<open>The field of rational numbers\<close>  | 
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
183  | 
|
| 
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
184  | 
context idom  | 
| 
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
185  | 
begin  | 
| 53196 | 186  | 
|
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
187  | 
subclass ring_no_zero_divisors ..  | 
| 53196 | 188  | 
|
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
189  | 
end  | 
| 
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
190  | 
|
| 
59867
 
58043346ca64
given up separate type classes demanding `inverse 0 = 0`
 
haftmann 
parents: 
58881 
diff
changeset
 | 
191  | 
instantiation fract :: (idom) field  | 
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
192  | 
begin  | 
| 
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
193  | 
|
| 61106 | 194  | 
lift_definition inverse_fract :: "'a fract \<Rightarrow> 'a fract"  | 
195  | 
is "\<lambda>x. if fst x = 0 then (0, 1) else (snd x, fst x)"  | 
|
196  | 
by(auto simp add: algebra_simps)  | 
|
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
197  | 
|
| 
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
198  | 
lemma inverse_fract [simp]: "inverse (Fract a b) = Fract (b::'a::idom) a"  | 
| 61106 | 199  | 
by transfer simp  | 
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
200  | 
|
| 
60429
 
d3d1e185cd63
uniform _ div _ as infix syntax for ring division
 
haftmann 
parents: 
60352 
diff
changeset
 | 
201  | 
definition divide_fract_def: "q div r = q * inverse (r:: 'a fract)"  | 
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
202  | 
|
| 
60429
 
d3d1e185cd63
uniform _ div _ as infix syntax for ring division
 
haftmann 
parents: 
60352 
diff
changeset
 | 
203  | 
lemma divide_fract [simp]: "Fract a b div Fract c d = Fract (a * d) (b * c)"  | 
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
204  | 
by (simp add: divide_fract_def)  | 
| 
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
205  | 
|
| 47252 | 206  | 
instance  | 
207  | 
proof  | 
|
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
208  | 
fix q :: "'a fract"  | 
| 
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
209  | 
assume "q \<noteq> 0"  | 
| 46573 | 210  | 
then show "inverse q * q = 1"  | 
211  | 
by (cases q rule: Fract_cases_nonzero)  | 
|
| 
57512
 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 
haftmann 
parents: 
54863 
diff
changeset
 | 
212  | 
(simp_all add: fract_expand eq_fract mult.commute)  | 
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
213  | 
next  | 
| 
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
214  | 
fix q r :: "'a fract"  | 
| 
60429
 
d3d1e185cd63
uniform _ div _ as infix syntax for ring division
 
haftmann 
parents: 
60352 
diff
changeset
 | 
215  | 
show "q div r = q * inverse r" by (simp add: divide_fract_def)  | 
| 36409 | 216  | 
next  | 
| 46573 | 217  | 
show "inverse 0 = (0:: 'a fract)"  | 
218  | 
by (simp add: fract_expand) (simp add: fract_collapse)  | 
|
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
219  | 
qed  | 
| 
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
220  | 
|
| 
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
221  | 
end  | 
| 
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
222  | 
|
| 
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents:  
diff
changeset
 | 
223  | 
|
| 60500 | 224  | 
subsubsection \<open>The ordered field of fractions over an ordered idom\<close>  | 
| 
36331
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
225  | 
|
| 61106 | 226  | 
instantiation fract :: (linordered_idom) linorder  | 
227  | 
begin  | 
|
| 
36331
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
228  | 
|
| 61106 | 229  | 
lemma less_eq_fract_respect:  | 
230  | 
fixes a b a' b' c d c' d' :: 'a  | 
|
231  | 
assumes neq: "b \<noteq> 0" "b' \<noteq> 0" "d \<noteq> 0" "d' \<noteq> 0"  | 
|
232  | 
assumes eq1: "a * b' = a' * b"  | 
|
233  | 
assumes eq2: "c * d' = c' * d"  | 
|
234  | 
shows "((a * d) * (b * d) \<le> (c * b) * (b * d)) \<longleftrightarrow> ((a' * d') * (b' * d') \<le> (c' * b') * (b' * d'))"  | 
|
235  | 
proof -  | 
|
| 
36331
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
236  | 
let ?le = "\<lambda>a b c d. ((a * d) * (b * d) \<le> (c * b) * (b * d))"  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
237  | 
  {
 | 
| 54463 | 238  | 
fix a b c d x :: 'a  | 
239  | 
assume x: "x \<noteq> 0"  | 
|
| 
36331
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
240  | 
have "?le a b c d = ?le (a * x) (b * x) c d"  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
241  | 
proof -  | 
| 54463 | 242  | 
from x have "0 < x * x"  | 
243  | 
by (auto simp add: zero_less_mult_iff)  | 
|
| 46573 | 244  | 
then have "?le a b c d =  | 
| 
36331
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
245  | 
((a * d) * (b * d) * (x * x) \<le> (c * b) * (b * d) * (x * x))"  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
246  | 
by (simp add: mult_le_cancel_right)  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
247  | 
also have "... = ?le (a * x) (b * x) c d"  | 
| 
57514
 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 
haftmann 
parents: 
57512 
diff
changeset
 | 
248  | 
by (simp add: ac_simps)  | 
| 
36331
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
249  | 
finally show ?thesis .  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
250  | 
qed  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
251  | 
} note le_factor = this  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
252  | 
|
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
253  | 
let ?D = "b * d" and ?D' = "b' * d'"  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
254  | 
from neq have D: "?D \<noteq> 0" by simp  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
255  | 
from neq have "?D' \<noteq> 0" by simp  | 
| 46573 | 256  | 
then have "?le a b c d = ?le (a * ?D') (b * ?D') c d"  | 
| 
36331
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
257  | 
by (rule le_factor)  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
258  | 
also have "... = ((a * b') * ?D * ?D' * d * d' \<le> (c * d') * ?D * ?D' * b * b')"  | 
| 
57514
 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 
haftmann 
parents: 
57512 
diff
changeset
 | 
259  | 
by (simp add: ac_simps)  | 
| 
36331
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
260  | 
also have "... = ((a' * b) * ?D * ?D' * d * d' \<le> (c' * d) * ?D * ?D' * b * b')"  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
261  | 
by (simp only: eq1 eq2)  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
262  | 
also have "... = ?le (a' * ?D) (b' * ?D) c' d'"  | 
| 
57514
 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 
haftmann 
parents: 
57512 
diff
changeset
 | 
263  | 
by (simp add: ac_simps)  | 
| 
36331
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
264  | 
also from D have "... = ?le a' b' c' d'"  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
265  | 
by (rule le_factor [symmetric])  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
266  | 
finally show "?le a b c d = ?le a' b' c' d'" .  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
267  | 
qed  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
268  | 
|
| 61106 | 269  | 
lift_definition less_eq_fract :: "'a fract \<Rightarrow> 'a fract \<Rightarrow> bool"  | 
270  | 
is "\<lambda>q r. (fst q * snd r) * (snd q * snd r) \<le> (fst r * snd q) * (snd q * snd r)"  | 
|
271  | 
by (clarsimp simp add: less_eq_fract_respect)  | 
|
| 
36331
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
272  | 
|
| 46573 | 273  | 
definition less_fract_def: "z < (w::'a fract) \<longleftrightarrow> z \<le> w \<and> \<not> w \<le> z"  | 
| 
36331
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
274  | 
|
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
275  | 
lemma le_fract [simp]:  | 
| 61106 | 276  | 
"\<lbrakk> b \<noteq> 0; d \<noteq> 0 \<rbrakk> \<Longrightarrow> Fract a b \<le> Fract c d \<longleftrightarrow> (a * d) * (b * d) \<le> (c * b) * (b * d)"  | 
277  | 
by transfer simp  | 
|
| 
36331
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
278  | 
|
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
279  | 
lemma less_fract [simp]:  | 
| 61106 | 280  | 
"\<lbrakk> b \<noteq> 0; d \<noteq> 0 \<rbrakk> \<Longrightarrow> Fract a b < Fract c d \<longleftrightarrow> (a * d) * (b * d) < (c * b) * (b * d)"  | 
| 63092 | 281  | 
by (simp add: less_fract_def less_le_not_le ac_simps)  | 
| 
36331
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
282  | 
|
| 47252 | 283  | 
instance  | 
284  | 
proof  | 
|
| 
36331
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
285  | 
fix q r s :: "'a fract"  | 
| 54463 | 286  | 
assume "q \<le> r" and "r \<le> s"  | 
287  | 
then show "q \<le> s"  | 
|
| 
36331
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
288  | 
proof (induct q, induct r, induct s)  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
289  | 
fix a b c d e f :: 'a  | 
| 54463 | 290  | 
assume neq: "b \<noteq> 0" "d \<noteq> 0" "f \<noteq> 0"  | 
291  | 
assume 1: "Fract a b \<le> Fract c d"  | 
|
292  | 
assume 2: "Fract c d \<le> Fract e f"  | 
|
| 
36331
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
293  | 
show "Fract a b \<le> Fract e f"  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
294  | 
proof -  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
295  | 
from neq obtain bb: "0 < b * b" and dd: "0 < d * d" and ff: "0 < f * f"  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
296  | 
by (auto simp add: zero_less_mult_iff linorder_neq_iff)  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
297  | 
have "(a * d) * (b * d) * (f * f) \<le> (c * b) * (b * d) * (f * f)"  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
298  | 
proof -  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
299  | 
from neq 1 have "(a * d) * (b * d) \<le> (c * b) * (b * d)"  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
300  | 
by simp  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
301  | 
with ff show ?thesis by (simp add: mult_le_cancel_right)  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
302  | 
qed  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
303  | 
also have "... = (c * f) * (d * f) * (b * b)"  | 
| 
57514
 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 
haftmann 
parents: 
57512 
diff
changeset
 | 
304  | 
by (simp only: ac_simps)  | 
| 
36331
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
305  | 
also have "... \<le> (e * d) * (d * f) * (b * b)"  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
306  | 
proof -  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
307  | 
from neq 2 have "(c * f) * (d * f) \<le> (e * d) * (d * f)"  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
308  | 
by simp  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
309  | 
with bb show ?thesis by (simp add: mult_le_cancel_right)  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
310  | 
qed  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
311  | 
finally have "(a * f) * (b * f) * (d * d) \<le> e * b * (b * f) * (d * d)"  | 
| 
57514
 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 
haftmann 
parents: 
57512 
diff
changeset
 | 
312  | 
by (simp only: ac_simps)  | 
| 
36331
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
313  | 
with dd have "(a * f) * (b * f) \<le> (e * b) * (b * f)"  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
314  | 
by (simp add: mult_le_cancel_right)  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
315  | 
with neq show ?thesis by simp  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
316  | 
qed  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
317  | 
qed  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
318  | 
next  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
319  | 
fix q r :: "'a fract"  | 
| 54463 | 320  | 
assume "q \<le> r" and "r \<le> q"  | 
321  | 
then show "q = r"  | 
|
| 
36331
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
322  | 
proof (induct q, induct r)  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
323  | 
fix a b c d :: 'a  | 
| 54463 | 324  | 
assume neq: "b \<noteq> 0" "d \<noteq> 0"  | 
325  | 
assume 1: "Fract a b \<le> Fract c d"  | 
|
326  | 
assume 2: "Fract c d \<le> Fract a b"  | 
|
| 
36331
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
327  | 
show "Fract a b = Fract c d"  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
328  | 
proof -  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
329  | 
from neq 1 have "(a * d) * (b * d) \<le> (c * b) * (b * d)"  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
330  | 
by simp  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
331  | 
also have "... \<le> (a * d) * (b * d)"  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
332  | 
proof -  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
333  | 
from neq 2 have "(c * b) * (d * b) \<le> (a * d) * (d * b)"  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
334  | 
by simp  | 
| 
57514
 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 
haftmann 
parents: 
57512 
diff
changeset
 | 
335  | 
then show ?thesis by (simp only: ac_simps)  | 
| 
36331
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
336  | 
qed  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
337  | 
finally have "(a * d) * (b * d) = (c * b) * (b * d)" .  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
338  | 
moreover from neq have "b * d \<noteq> 0" by simp  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
339  | 
ultimately have "a * d = c * b" by simp  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
340  | 
with neq show ?thesis by (simp add: eq_fract)  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
341  | 
qed  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
342  | 
qed  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
343  | 
next  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
344  | 
fix q r :: "'a fract"  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
345  | 
show "q \<le> q"  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
346  | 
by (induct q) simp  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
347  | 
show "(q < r) = (q \<le> r \<and> \<not> r \<le> q)"  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
348  | 
by (simp only: less_fract_def)  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
349  | 
show "q \<le> r \<or> r \<le> q"  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
350  | 
by (induct q, induct r)  | 
| 
57512
 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 
haftmann 
parents: 
54863 
diff
changeset
 | 
351  | 
(simp add: mult.commute, rule linorder_linear)  | 
| 
36331
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
352  | 
qed  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
353  | 
|
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
354  | 
end  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
355  | 
|
| 54463 | 356  | 
instantiation fract :: (linordered_idom) "{distrib_lattice,abs_if,sgn_if}"
 | 
| 
36331
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
357  | 
begin  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
358  | 
|
| 61106 | 359  | 
definition abs_fract_def2: "\<bar>q\<bar> = (if q < 0 then -q else (q::'a fract))"  | 
| 
36331
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
360  | 
|
| 46573 | 361  | 
definition sgn_fract_def:  | 
| 54463 | 362  | 
"sgn (q::'a fract) = (if q = 0 then 0 else if 0 < q then 1 else - 1)"  | 
| 
36331
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
363  | 
|
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
364  | 
theorem abs_fract [simp]: "\<bar>Fract a b\<bar> = Fract \<bar>a\<bar> \<bar>b\<bar>"  | 
| 61106 | 365  | 
unfolding abs_fract_def2 not_le[symmetric]  | 
366  | 
by transfer(auto simp add: zero_less_mult_iff le_less)  | 
|
| 
36331
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
367  | 
|
| 46573 | 368  | 
definition inf_fract_def:  | 
| 61076 | 369  | 
"(inf :: 'a fract \<Rightarrow> 'a fract \<Rightarrow> 'a fract) = min"  | 
| 
36331
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
370  | 
|
| 46573 | 371  | 
definition sup_fract_def:  | 
| 61076 | 372  | 
"(sup :: 'a fract \<Rightarrow> 'a fract \<Rightarrow> 'a fract) = max"  | 
| 
36331
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
373  | 
|
| 46573 | 374  | 
instance  | 
| 61106 | 375  | 
by intro_classes (simp_all add: abs_fract_def2 sgn_fract_def inf_fract_def sup_fract_def max_min_distrib2)  | 
| 
36331
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
376  | 
|
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
377  | 
end  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
378  | 
|
| 
59867
 
58043346ca64
given up separate type classes demanding `inverse 0 = 0`
 
haftmann 
parents: 
58881 
diff
changeset
 | 
379  | 
instance fract :: (linordered_idom) linordered_field  | 
| 
36331
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
380  | 
proof  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
381  | 
fix q r s :: "'a fract"  | 
| 53196 | 382  | 
assume "q \<le> r"  | 
383  | 
then show "s + q \<le> s + r"  | 
|
| 
36331
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
384  | 
proof (induct q, induct r, induct s)  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
385  | 
fix a b c d e f :: 'a  | 
| 53196 | 386  | 
assume neq: "b \<noteq> 0" "d \<noteq> 0" "f \<noteq> 0"  | 
| 
36331
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
387  | 
assume le: "Fract a b \<le> Fract c d"  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
388  | 
show "Fract e f + Fract a b \<le> Fract e f + Fract c d"  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
389  | 
proof -  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
390  | 
let ?F = "f * f" from neq have F: "0 < ?F"  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
391  | 
by (auto simp add: zero_less_mult_iff)  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
392  | 
from neq le have "(a * d) * (b * d) \<le> (c * b) * (b * d)"  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
393  | 
by simp  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
394  | 
with F have "(a * d) * (b * d) * ?F * ?F \<le> (c * b) * (b * d) * ?F * ?F"  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
395  | 
by (simp add: mult_le_cancel_right)  | 
| 
36348
 
89c54f51f55a
dropped group_simps, ring_simps, field_eq_simps; classes division_ring_inverse_zero, field_inverse_zero, linordered_field_inverse_zero
 
haftmann 
parents: 
36331 
diff
changeset
 | 
396  | 
with neq show ?thesis by (simp add: field_simps)  | 
| 
36331
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
397  | 
qed  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
398  | 
qed  | 
| 53196 | 399  | 
next  | 
400  | 
fix q r s :: "'a fract"  | 
|
401  | 
assume "q < r" and "0 < s"  | 
|
402  | 
then show "s * q < s * r"  | 
|
| 
36331
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
403  | 
proof (induct q, induct r, induct s)  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
404  | 
fix a b c d e f :: 'a  | 
| 54463 | 405  | 
assume neq: "b \<noteq> 0" "d \<noteq> 0" "f \<noteq> 0"  | 
| 
36331
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
406  | 
assume le: "Fract a b < Fract c d"  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
407  | 
assume gt: "0 < Fract e f"  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
408  | 
show "Fract e f * Fract a b < Fract e f * Fract c d"  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
409  | 
proof -  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
410  | 
let ?E = "e * f" and ?F = "f * f"  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
411  | 
from neq gt have "0 < ?E"  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
412  | 
by (auto simp add: Zero_fract_def order_less_le eq_fract)  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
413  | 
moreover from neq have "0 < ?F"  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
414  | 
by (auto simp add: zero_less_mult_iff)  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
415  | 
moreover from neq le have "(a * d) * (b * d) < (c * b) * (b * d)"  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
416  | 
by simp  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
417  | 
ultimately have "(a * d) * (b * d) * ?E * ?F < (c * b) * (b * d) * ?E * ?F"  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
418  | 
by (simp add: mult_less_cancel_right)  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
419  | 
with neq show ?thesis  | 
| 
57514
 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 
haftmann 
parents: 
57512 
diff
changeset
 | 
420  | 
by (simp add: ac_simps)  | 
| 
36331
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
421  | 
qed  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
422  | 
qed  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
423  | 
qed  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
424  | 
|
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
425  | 
lemma fract_induct_pos [case_names Fract]:  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
426  | 
fixes P :: "'a::linordered_idom fract \<Rightarrow> bool"  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
427  | 
assumes step: "\<And>a b. 0 < b \<Longrightarrow> P (Fract a b)"  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
428  | 
shows "P q"  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
429  | 
proof (cases q)  | 
| 54463 | 430  | 
case (Fract a b)  | 
431  | 
  {
 | 
|
432  | 
fix a b :: 'a  | 
|
| 
36331
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
433  | 
assume b: "b < 0"  | 
| 54463 | 434  | 
have "P (Fract a b)"  | 
435  | 
proof -  | 
|
436  | 
from b have "0 < - b" by simp  | 
|
437  | 
then have "P (Fract (- a) (- b))"  | 
|
438  | 
by (rule step)  | 
|
439  | 
then show "P (Fract a b)"  | 
|
440  | 
by (simp add: order_less_imp_not_eq [OF b])  | 
|
441  | 
qed  | 
|
442  | 
}  | 
|
443  | 
with Fract show "P q"  | 
|
444  | 
by (auto simp add: linorder_neq_iff step)  | 
|
| 
36331
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
445  | 
qed  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
446  | 
|
| 53196 | 447  | 
lemma zero_less_Fract_iff: "0 < b \<Longrightarrow> 0 < Fract a b \<longleftrightarrow> 0 < a"  | 
| 
36331
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
448  | 
by (auto simp add: Zero_fract_def zero_less_mult_iff)  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
449  | 
|
| 53196 | 450  | 
lemma Fract_less_zero_iff: "0 < b \<Longrightarrow> Fract a b < 0 \<longleftrightarrow> a < 0"  | 
| 
36331
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
451  | 
by (auto simp add: Zero_fract_def mult_less_0_iff)  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
452  | 
|
| 53196 | 453  | 
lemma zero_le_Fract_iff: "0 < b \<Longrightarrow> 0 \<le> Fract a b \<longleftrightarrow> 0 \<le> a"  | 
| 
36331
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
454  | 
by (auto simp add: Zero_fract_def zero_le_mult_iff)  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
455  | 
|
| 53196 | 456  | 
lemma Fract_le_zero_iff: "0 < b \<Longrightarrow> Fract a b \<le> 0 \<longleftrightarrow> a \<le> 0"  | 
| 
36331
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
457  | 
by (auto simp add: Zero_fract_def mult_le_0_iff)  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
458  | 
|
| 53196 | 459  | 
lemma one_less_Fract_iff: "0 < b \<Longrightarrow> 1 < Fract a b \<longleftrightarrow> b < a"  | 
| 
36331
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
460  | 
by (auto simp add: One_fract_def mult_less_cancel_right_disj)  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
461  | 
|
| 53196 | 462  | 
lemma Fract_less_one_iff: "0 < b \<Longrightarrow> Fract a b < 1 \<longleftrightarrow> a < b"  | 
| 
36331
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
463  | 
by (auto simp add: One_fract_def mult_less_cancel_right_disj)  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
464  | 
|
| 53196 | 465  | 
lemma one_le_Fract_iff: "0 < b \<Longrightarrow> 1 \<le> Fract a b \<longleftrightarrow> b \<le> a"  | 
| 
36331
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
466  | 
by (auto simp add: One_fract_def mult_le_cancel_right)  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
467  | 
|
| 53196 | 468  | 
lemma Fract_le_one_iff: "0 < b \<Longrightarrow> Fract a b \<le> 1 \<longleftrightarrow> a \<le> b"  | 
| 
36331
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
469  | 
by (auto simp add: One_fract_def mult_le_cancel_right)  | 
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
470  | 
|
| 
 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 
huffman 
parents: 
36312 
diff
changeset
 | 
471  | 
end  |