| author | wenzelm | 
| Sat, 05 Mar 2022 21:52:21 +0100 | |
| changeset 75228 | 33fb3014876f | 
| parent 74362 | 0135a0c77b64 | 
| child 76121 | f58ad163bb75 | 
| permissions | -rw-r--r-- | 
| 65435 | 1 | (* Title: HOL/Computational_Algebra/Normalized_Fraction.thy | 
| 64591 
240a39af9ec4
restructured matter on polynomials and normalized fractions
 haftmann parents: 
64240diff
changeset | 2 | Author: Manuel Eberl | 
| 
240a39af9ec4
restructured matter on polynomials and normalized fractions
 haftmann parents: 
64240diff
changeset | 3 | *) | 
| 
240a39af9ec4
restructured matter on polynomials and normalized fractions
 haftmann parents: 
64240diff
changeset | 4 | |
| 63500 | 5 | theory Normalized_Fraction | 
| 6 | imports | |
| 7 | Main | |
| 65417 | 8 | Euclidean_Algorithm | 
| 65366 | 9 | Fraction_Field | 
| 63500 | 10 | begin | 
| 11 | ||
| 12 | definition quot_to_fract :: "'a :: {idom} \<times> 'a \<Rightarrow> 'a fract" where
 | |
| 13 | "quot_to_fract = (\<lambda>(a,b). Fraction_Field.Fract a b)" | |
| 14 | ||
| 71398 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 Manuel Eberl <eberlm@in.tum.de> parents: 
67078diff
changeset | 15 | definition normalize_quot :: "'a :: {ring_gcd,idom_divide,semiring_gcd_mult_normalize} \<times> 'a \<Rightarrow> 'a \<times> 'a" where
 | 
| 63500 | 16 | "normalize_quot = | 
| 17 | (\<lambda>(a,b). if b = 0 then (0,1) else let d = gcd a b * unit_factor b in (a div d, b div d))" | |
| 18 | ||
| 67051 | 19 | lemma normalize_quot_zero [simp]: | 
| 20 | "normalize_quot (a, 0) = (0, 1)" | |
| 21 | by (simp add: normalize_quot_def) | |
| 22 | ||
| 67078 
6a85b8a9c28c
removed overambitious simp rules from e7e54a0b9197
 haftmann parents: 
67051diff
changeset | 23 | lemma normalize_quot_proj: | 
| 67051 | 24 | "fst (normalize_quot (a, b)) = a div (gcd a b * unit_factor b)" | 
| 25 | "snd (normalize_quot (a, b)) = normalize b div gcd a b" if "b \<noteq> 0" | |
| 26 | using that by (simp_all add: normalize_quot_def Let_def mult.commute [of _ "unit_factor b"] dvd_div_mult2_eq mult_unit_dvd_iff') | |
| 27 | ||
| 63500 | 28 | definition normalized_fracts :: "('a :: {ring_gcd,idom_divide} \<times> 'a) set" where
 | 
| 29 |   "normalized_fracts = {(a,b). coprime a b \<and> unit_factor b = 1}"
 | |
| 30 | ||
| 31 | lemma not_normalized_fracts_0_denom [simp]: "(a, 0) \<notin> normalized_fracts" | |
| 32 | by (auto simp: normalized_fracts_def) | |
| 33 | ||
| 34 | lemma unit_factor_snd_normalize_quot [simp]: | |
| 35 | "unit_factor (snd (normalize_quot x)) = 1" | |
| 36 | by (simp add: normalize_quot_def case_prod_unfold Let_def dvd_unit_factor_div | |
| 37 | mult_unit_dvd_iff unit_factor_mult unit_factor_gcd) | |
| 38 | ||
| 39 | lemma snd_normalize_quot_nonzero [simp]: "snd (normalize_quot x) \<noteq> 0" | |
| 40 | using unit_factor_snd_normalize_quot[of x] | |
| 41 | by (auto simp del: unit_factor_snd_normalize_quot) | |
| 42 | ||
| 43 | lemma normalize_quot_aux: | |
| 44 | fixes a b | |
| 45 | assumes "b \<noteq> 0" | |
| 46 | defines "d \<equiv> gcd a b * unit_factor b" | |
| 47 | shows "a = fst (normalize_quot (a,b)) * d" "b = snd (normalize_quot (a,b)) * d" | |
| 48 | "d dvd a" "d dvd b" "d \<noteq> 0" | |
| 49 | proof - | |
| 50 | from assms show "d dvd a" "d dvd b" | |
| 51 | by (simp_all add: d_def mult_unit_dvd_iff) | |
| 52 | thus "a = fst (normalize_quot (a,b)) * d" "b = snd (normalize_quot (a,b)) * d" "d \<noteq> 0" | |
| 53 | by (auto simp: normalize_quot_def Let_def d_def \<open>b \<noteq> 0\<close>) | |
| 54 | qed | |
| 55 | ||
| 56 | lemma normalize_quotE: | |
| 57 | assumes "b \<noteq> 0" | |
| 58 | obtains d where "a = fst (normalize_quot (a,b)) * d" "b = snd (normalize_quot (a,b)) * d" | |
| 59 | "d dvd a" "d dvd b" "d \<noteq> 0" | |
| 60 | using that[OF normalize_quot_aux[OF assms]] . | |
| 61 | ||
| 62 | lemma normalize_quotE': | |
| 63 | assumes "snd x \<noteq> 0" | |
| 64 | obtains d where "fst x = fst (normalize_quot x) * d" "snd x = snd (normalize_quot x) * d" | |
| 65 | "d dvd fst x" "d dvd snd x" "d \<noteq> 0" | |
| 66 | proof - | |
| 74362 | 67 | from normalize_quotE[OF assms, of "fst x"] obtain d where | 
| 68 | "fst x = fst (normalize_quot (fst x, snd x)) * d" | |
| 69 | "snd x = snd (normalize_quot (fst x, snd x)) * d" | |
| 70 | "d dvd fst x" | |
| 71 | "d dvd snd x" | |
| 72 | "d \<noteq> 0" . | |
| 73 | then show ?thesis unfolding prod.collapse by (intro that[of d]) | |
| 63500 | 74 | qed | 
| 75 | ||
| 76 | lemma coprime_normalize_quot: | |
| 77 | "coprime (fst (normalize_quot x)) (snd (normalize_quot x))" | |
| 67051 | 78 | by (simp add: normalize_quot_def case_prod_unfold div_mult_unit2) | 
| 79 | (metis coprime_mult_self_right_iff div_gcd_coprime unit_div_mult_self unit_factor_is_unit) | |
| 63500 | 80 | |
| 81 | lemma normalize_quot_in_normalized_fracts [simp]: "normalize_quot x \<in> normalized_fracts" | |
| 82 | by (simp add: normalized_fracts_def coprime_normalize_quot case_prod_unfold) | |
| 83 | ||
| 84 | lemma normalize_quot_eq_iff: | |
| 85 | assumes "b \<noteq> 0" "d \<noteq> 0" | |
| 86 | shows "normalize_quot (a,b) = normalize_quot (c,d) \<longleftrightarrow> a * d = b * c" | |
| 87 | proof - | |
| 88 | define x y where "x = normalize_quot (a,b)" and "y = normalize_quot (c,d)" | |
| 89 | from normalize_quotE[OF assms(1), of a] normalize_quotE[OF assms(2), of c] | |
| 90 | obtain d1 d2 | |
| 91 | where "a = fst x * d1" "b = snd x * d1" "c = fst y * d2" "d = snd y * d2" "d1 \<noteq> 0" "d2 \<noteq> 0" | |
| 92 | unfolding x_def y_def by metis | |
| 93 | hence "a * d = b * c \<longleftrightarrow> fst x * snd y = snd x * fst y" by simp | |
| 94 | also have "\<dots> \<longleftrightarrow> fst x = fst y \<and> snd x = snd y" | |
| 95 | by (intro coprime_crossproduct') (simp_all add: x_def y_def coprime_normalize_quot) | |
| 96 | also have "\<dots> \<longleftrightarrow> x = y" using prod_eqI by blast | |
| 97 | finally show "x = y \<longleftrightarrow> a * d = b * c" .. | |
| 98 | qed | |
| 99 | ||
| 100 | lemma normalize_quot_eq_iff': | |
| 101 | assumes "snd x \<noteq> 0" "snd y \<noteq> 0" | |
| 102 | shows "normalize_quot x = normalize_quot y \<longleftrightarrow> fst x * snd y = snd x * fst y" | |
| 103 | using assms by (cases x, cases y, hypsubst) (subst normalize_quot_eq_iff, simp_all) | |
| 104 | ||
| 105 | lemma normalize_quot_id: "x \<in> normalized_fracts \<Longrightarrow> normalize_quot x = x" | |
| 106 | by (auto simp: normalized_fracts_def normalize_quot_def case_prod_unfold) | |
| 107 | ||
| 108 | lemma normalize_quot_idem [simp]: "normalize_quot (normalize_quot x) = normalize_quot x" | |
| 109 | by (rule normalize_quot_id) simp_all | |
| 110 | ||
| 111 | lemma fractrel_iff_normalize_quot_eq: | |
| 112 | "fractrel x y \<longleftrightarrow> normalize_quot x = normalize_quot y \<and> snd x \<noteq> 0 \<and> snd y \<noteq> 0" | |
| 113 | by (cases x, cases y) (auto simp: fractrel_def normalize_quot_eq_iff) | |
| 114 | ||
| 115 | lemma fractrel_normalize_quot_left: | |
| 116 | assumes "snd x \<noteq> 0" | |
| 117 | shows "fractrel (normalize_quot x) y \<longleftrightarrow> fractrel x y" | |
| 118 | using assms by (subst (1 2) fractrel_iff_normalize_quot_eq) auto | |
| 119 | ||
| 120 | lemma fractrel_normalize_quot_right: | |
| 121 | assumes "snd x \<noteq> 0" | |
| 122 | shows "fractrel y (normalize_quot x) \<longleftrightarrow> fractrel y x" | |
| 123 | using assms by (subst (1 2) fractrel_iff_normalize_quot_eq) auto | |
| 124 | ||
| 125 | ||
| 71398 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 Manuel Eberl <eberlm@in.tum.de> parents: 
67078diff
changeset | 126 | lift_definition quot_of_fract :: | 
| 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 Manuel Eberl <eberlm@in.tum.de> parents: 
67078diff
changeset | 127 |   "'a :: {ring_gcd,idom_divide,semiring_gcd_mult_normalize} fract \<Rightarrow> 'a \<times> 'a" 
 | 
| 63500 | 128 | is normalize_quot | 
| 129 | by (subst (asm) fractrel_iff_normalize_quot_eq) simp_all | |
| 130 | ||
| 131 | lemma quot_to_fract_quot_of_fract [simp]: "quot_to_fract (quot_of_fract x) = x" | |
| 132 | unfolding quot_to_fract_def | |
| 133 | proof transfer | |
| 134 | fix x :: "'a \<times> 'a" assume rel: "fractrel x x" | |
| 135 | define x' where "x' = normalize_quot x" | |
| 136 | obtain a b where [simp]: "x = (a, b)" by (cases x) | |
| 137 | from rel have "b \<noteq> 0" by simp | |
| 74362 | 138 | from normalize_quotE[OF this, of a] obtain d | 
| 139 | where | |
| 140 | "a = fst (normalize_quot (a, b)) * d" | |
| 141 | "b = snd (normalize_quot (a, b)) * d" | |
| 142 | "d dvd a" | |
| 143 | "d dvd b" | |
| 144 | "d \<noteq> 0" . | |
| 63500 | 145 | hence "a = fst x' * d" "b = snd x' * d" "d \<noteq> 0" "snd x' \<noteq> 0" by (simp_all add: x'_def) | 
| 146 | thus "fractrel (case x' of (a, b) \<Rightarrow> if b = 0 then (0, 1) else (a, b)) x" | |
| 147 | by (auto simp add: case_prod_unfold) | |
| 148 | qed | |
| 149 | ||
| 150 | lemma quot_of_fract_quot_to_fract: "quot_of_fract (quot_to_fract x) = normalize_quot x" | |
| 151 | proof (cases "snd x = 0") | |
| 152 | case True | |
| 153 | thus ?thesis unfolding quot_to_fract_def | |
| 154 | by transfer (simp add: case_prod_unfold normalize_quot_def) | |
| 155 | next | |
| 156 | case False | |
| 157 | thus ?thesis unfolding quot_to_fract_def by transfer (simp add: case_prod_unfold) | |
| 158 | qed | |
| 159 | ||
| 160 | lemma quot_of_fract_quot_to_fract': | |
| 161 | "x \<in> normalized_fracts \<Longrightarrow> quot_of_fract (quot_to_fract x) = x" | |
| 162 | unfolding quot_to_fract_def by transfer (auto simp: normalize_quot_id) | |
| 163 | ||
| 164 | lemma quot_of_fract_in_normalized_fracts [simp]: "quot_of_fract x \<in> normalized_fracts" | |
| 165 | by transfer simp | |
| 166 | ||
| 167 | lemma normalize_quotI: | |
| 168 | assumes "a * d = b * c" "b \<noteq> 0" "(c, d) \<in> normalized_fracts" | |
| 169 | shows "normalize_quot (a, b) = (c, d)" | |
| 170 | proof - | |
| 171 | from assms have "normalize_quot (a, b) = normalize_quot (c, d)" | |
| 172 | by (subst normalize_quot_eq_iff) auto | |
| 173 | also have "\<dots> = (c, d)" by (intro normalize_quot_id) fact | |
| 174 | finally show ?thesis . | |
| 175 | qed | |
| 176 | ||
| 177 | lemma td_normalized_fract: | |
| 178 | "type_definition quot_of_fract quot_to_fract normalized_fracts" | |
| 179 | by standard (simp_all add: quot_of_fract_quot_to_fract') | |
| 180 | ||
| 181 | lemma quot_of_fract_add_aux: | |
| 182 | assumes "snd x \<noteq> 0" "snd y \<noteq> 0" | |
| 183 | shows "(fst x * snd y + fst y * snd x) * (snd (normalize_quot x) * snd (normalize_quot y)) = | |
| 184 | snd x * snd y * (fst (normalize_quot x) * snd (normalize_quot y) + | |
| 185 | snd (normalize_quot x) * fst (normalize_quot y))" | |
| 186 | proof - | |
| 74362 | 187 | from normalize_quotE'[OF assms(1)] obtain d | 
| 188 | where d: | |
| 189 | "fst x = fst (normalize_quot x) * d" | |
| 190 | "snd x = snd (normalize_quot x) * d" | |
| 191 | "d dvd fst x" | |
| 192 | "d dvd snd x" | |
| 193 | "d \<noteq> 0" . | |
| 194 | from normalize_quotE'[OF assms(2)] obtain e | |
| 195 | where e: | |
| 196 | "fst y = fst (normalize_quot y) * e" | |
| 197 | "snd y = snd (normalize_quot y) * e" | |
| 198 | "e dvd fst y" | |
| 199 | "e dvd snd y" | |
| 200 | "e \<noteq> 0" . | |
| 63500 | 201 | show ?thesis by (simp_all add: d e algebra_simps) | 
| 202 | qed | |
| 203 | ||
| 204 | ||
| 205 | locale fract_as_normalized_quot | |
| 206 | begin | |
| 207 | setup_lifting td_normalized_fract | |
| 208 | end | |
| 209 | ||
| 210 | ||
| 211 | lemma quot_of_fract_add: | |
| 212 | "quot_of_fract (x + y) = | |
| 213 | (let (a,b) = quot_of_fract x; (c,d) = quot_of_fract y | |
| 214 | in normalize_quot (a * d + b * c, b * d))" | |
| 215 | by transfer (insert quot_of_fract_add_aux, | |
| 216 | simp_all add: Let_def case_prod_unfold normalize_quot_eq_iff) | |
| 217 | ||
| 218 | lemma quot_of_fract_uminus: | |
| 219 | "quot_of_fract (-x) = (let (a,b) = quot_of_fract x in (-a, b))" | |
| 64592 
7759f1766189
more fine-grained type class hierarchy for div and mod
 haftmann parents: 
64591diff
changeset | 220 | by transfer (auto simp: case_prod_unfold Let_def normalize_quot_def dvd_neg_div mult_unit_dvd_iff) | 
| 63500 | 221 | |
| 222 | lemma quot_of_fract_diff: | |
| 223 | "quot_of_fract (x - y) = | |
| 224 | (let (a,b) = quot_of_fract x; (c,d) = quot_of_fract y | |
| 225 | in normalize_quot (a * d - b * c, b * d))" (is "_ = ?rhs") | |
| 226 | proof - | |
| 227 | have "x - y = x + -y" by simp | |
| 228 | also have "quot_of_fract \<dots> = ?rhs" | |
| 229 | by (simp only: quot_of_fract_add quot_of_fract_uminus Let_def case_prod_unfold) simp_all | |
| 230 | finally show ?thesis . | |
| 231 | qed | |
| 232 | ||
| 233 | lemma normalize_quot_mult_coprime: | |
| 234 | assumes "coprime a b" "coprime c d" "unit_factor b = 1" "unit_factor d = 1" | |
| 235 | defines "e \<equiv> fst (normalize_quot (a, d))" and "f \<equiv> snd (normalize_quot (a, d))" | |
| 236 | and "g \<equiv> fst (normalize_quot (c, b))" and "h \<equiv> snd (normalize_quot (c, b))" | |
| 237 | shows "normalize_quot (a * c, b * d) = (e * g, f * h)" | |
| 238 | proof (rule normalize_quotI) | |
| 67051 | 239 | from assms have "gcd a b = 1" "gcd c d = 1" | 
| 240 | by simp_all | |
| 63500 | 241 | from assms have "b \<noteq> 0" "d \<noteq> 0" by auto | 
| 67051 | 242 | with assms have "normalize b = b" "normalize d = d" | 
| 243 | by (auto intro: normalize_unit_factor_eqI) | |
| 74362 | 244 | from normalize_quotE [OF \<open>b \<noteq> 0\<close>, of c] obtain k | 
| 245 | where | |
| 246 | "c = fst (normalize_quot (c, b)) * k" | |
| 247 | "b = snd (normalize_quot (c, b)) * k" | |
| 248 | "k dvd c" "k dvd b" "k \<noteq> 0" . | |
| 67051 | 249 | note k = this [folded \<open>gcd a b = 1\<close> \<open>gcd c d = 1\<close> assms(3) assms(4)] | 
| 74362 | 250 | from normalize_quotE [OF \<open>d \<noteq> 0\<close>, of a] obtain l | 
| 251 | where "a = fst (normalize_quot (a, d)) * l" | |
| 252 | "d = snd (normalize_quot (a, d)) * l" | |
| 253 | "l dvd a" "l dvd d" "l \<noteq> 0" . | |
| 67051 | 254 | note l = this [folded \<open>gcd a b = 1\<close> \<open>gcd c d = 1\<close> assms(3) assms(4)] | 
| 255 | from k l show "a * c * (f * h) = b * d * (e * g)" | |
| 256 | by (metis e_def f_def g_def h_def mult.commute mult.left_commute) | |
| 63500 | 257 | from assms have [simp]: "unit_factor f = 1" "unit_factor h = 1" | 
| 258 | by simp_all | |
| 259 | from assms have "coprime e f" "coprime g h" by (simp_all add: coprime_normalize_quot) | |
| 67051 | 260 | with k l assms(1,2) \<open>b \<noteq> 0\<close> \<open>d \<noteq> 0\<close> \<open>unit_factor b = 1\<close> \<open>unit_factor d = 1\<close> | 
| 261 | \<open>normalize b = b\<close> \<open>normalize d = d\<close> | |
| 262 | show "(e * g, f * h) \<in> normalized_fracts" | |
| 263 | by (simp add: normalized_fracts_def unit_factor_mult e_def f_def g_def h_def | |
| 264 | coprime_normalize_quot dvd_unit_factor_div unit_factor_gcd) | |
| 265 | (metis coprime_mult_left_iff coprime_mult_right_iff) | |
| 63500 | 266 | qed (insert assms(3,4), auto) | 
| 267 | ||
| 268 | lemma normalize_quot_mult: | |
| 269 | assumes "snd x \<noteq> 0" "snd y \<noteq> 0" | |
| 270 | shows "normalize_quot (fst x * fst y, snd x * snd y) = normalize_quot | |
| 271 | (fst (normalize_quot x) * fst (normalize_quot y), | |
| 272 | snd (normalize_quot x) * snd (normalize_quot y))" | |
| 273 | proof - | |
| 74362 | 274 | from normalize_quotE'[OF assms(1)] obtain d where d: | 
| 275 | "fst x = fst (normalize_quot x) * d" | |
| 276 | "snd x = snd (normalize_quot x) * d" | |
| 277 | "d dvd fst x" | |
| 278 | "d dvd snd x" | |
| 279 | "d \<noteq> 0" . | |
| 280 | from normalize_quotE'[OF assms(2)] obtain e where e: | |
| 281 | "fst y = fst (normalize_quot y) * e" | |
| 282 | "snd y = snd (normalize_quot y) * e" | |
| 283 | "e dvd fst y" | |
| 284 | "e dvd snd y" | |
| 285 | "e \<noteq> 0" . | |
| 63500 | 286 | show ?thesis by (simp_all add: d e algebra_simps normalize_quot_eq_iff) | 
| 287 | qed | |
| 288 | ||
| 289 | lemma quot_of_fract_mult: | |
| 290 | "quot_of_fract (x * y) = | |
| 291 | (let (a,b) = quot_of_fract x; (c,d) = quot_of_fract y; | |
| 292 | (e,f) = normalize_quot (a,d); (g,h) = normalize_quot (c,b) | |
| 293 | in (e*g, f*h))" | |
| 67051 | 294 | by transfer | 
| 295 | (simp add: split_def Let_def coprime_normalize_quot normalize_quot_mult normalize_quot_mult_coprime) | |
| 63500 | 296 | |
| 297 | lemma normalize_quot_0 [simp]: | |
| 298 | "normalize_quot (0, x) = (0, 1)" "normalize_quot (x, 0) = (0, 1)" | |
| 299 | by (simp_all add: normalize_quot_def) | |
| 300 | ||
| 301 | lemma normalize_quot_eq_0_iff [simp]: "fst (normalize_quot x) = 0 \<longleftrightarrow> fst x = 0 \<or> snd x = 0" | |
| 302 | by (auto simp: normalize_quot_def case_prod_unfold Let_def div_mult_unit2 dvd_div_eq_0_iff) | |
| 303 | ||
| 304 | lemma fst_quot_of_fract_0_imp: "fst (quot_of_fract x) = 0 \<Longrightarrow> snd (quot_of_fract x) = 1" | |
| 305 | by transfer auto | |
| 306 | ||
| 307 | lemma normalize_quot_swap: | |
| 308 | assumes "a \<noteq> 0" "b \<noteq> 0" | |
| 309 | defines "a' \<equiv> fst (normalize_quot (a, b))" and "b' \<equiv> snd (normalize_quot (a, b))" | |
| 310 | shows "normalize_quot (b, a) = (b' div unit_factor a', a' div unit_factor a')" | |
| 311 | proof (rule normalize_quotI) | |
| 74362 | 312 | from normalize_quotE[OF assms(2), of a] obtain d where | 
| 313 | "a = fst (normalize_quot (a, b)) * d" | |
| 314 | "b = snd (normalize_quot (a, b)) * d" | |
| 315 | "d dvd a" "d dvd b" "d \<noteq> 0" . | |
| 316 | note d = this [folded assms(3,4)] | |
| 63500 | 317 | show "b * (a' div unit_factor a') = a * (b' div unit_factor a')" | 
| 318 | using assms(1,2) d | |
| 319 | by (simp add: div_unit_factor [symmetric] unit_div_mult_swap mult_ac del: div_unit_factor) | |
| 320 | have "coprime a' b'" by (simp add: a'_def b'_def coprime_normalize_quot) | |
| 321 | thus "(b' div unit_factor a', a' div unit_factor a') \<in> normalized_fracts" | |
| 67051 | 322 | using assms(1,2) d | 
| 67078 
6a85b8a9c28c
removed overambitious simp rules from e7e54a0b9197
 haftmann parents: 
67051diff
changeset | 323 | by (auto simp add: normalized_fracts_def ac_simps dvd_div_unit_iff elim: coprime_imp_coprime) | 
| 63500 | 324 | qed fact+ | 
| 325 | ||
| 326 | lemma quot_of_fract_inverse: | |
| 327 | "quot_of_fract (inverse x) = | |
| 328 | (let (a,b) = quot_of_fract x; d = unit_factor a | |
| 329 | in if d = 0 then (0, 1) else (b div d, a div d))" | |
| 330 | proof (transfer, goal_cases) | |
| 331 | case (1 x) | |
| 332 | from normalize_quot_swap[of "fst x" "snd x"] show ?case | |
| 333 | by (auto simp: Let_def case_prod_unfold) | |
| 334 | qed | |
| 335 | ||
| 336 | lemma normalize_quot_div_unit_left: | |
| 337 | fixes x y u | |
| 338 | assumes "is_unit u" | |
| 339 | defines "x' \<equiv> fst (normalize_quot (x, y))" and "y' \<equiv> snd (normalize_quot (x, y))" | |
| 340 | shows "normalize_quot (x div u, y) = (x' div u, y')" | |
| 341 | proof (cases "y = 0") | |
| 342 | case False | |
| 67051 | 343 | define v where "v = 1 div u" | 
| 344 | with \<open>is_unit u\<close> have "is_unit v" and u: "\<And>a. a div u = a * v" | |
| 345 | by simp_all | |
| 346 | from \<open>is_unit v\<close> have "coprime v = top" | |
| 347 | by (simp add: fun_eq_iff is_unit_left_imp_coprime) | |
| 74362 | 348 | from normalize_quotE[OF False, of x] obtain d where | 
| 349 | "x = fst (normalize_quot (x, y)) * d" | |
| 350 | "y = snd (normalize_quot (x, y)) * d" | |
| 351 | "d dvd x" "d dvd y" "d \<noteq> 0" . | |
| 67051 | 352 | note d = this[folded assms(2,3)] | 
| 353 | from assms have "coprime x' y'" "unit_factor y' = 1" | |
| 354 | by (simp_all add: coprime_normalize_quot) | |
| 355 | with d \<open>coprime v = top\<close> have "normalize_quot (x * v, y) = (x' * v, y')" | |
| 356 | by (auto simp: normalized_fracts_def intro: normalize_quotI) | |
| 357 | then show ?thesis | |
| 358 | by (simp add: u) | |
| 63500 | 359 | qed (simp_all add: assms) | 
| 360 | ||
| 361 | lemma normalize_quot_div_unit_right: | |
| 362 | fixes x y u | |
| 363 | assumes "is_unit u" | |
| 364 | defines "x' \<equiv> fst (normalize_quot (x, y))" and "y' \<equiv> snd (normalize_quot (x, y))" | |
| 365 | shows "normalize_quot (x, y div u) = (x' * u, y')" | |
| 366 | proof (cases "y = 0") | |
| 367 | case False | |
| 74362 | 368 | from normalize_quotE[OF this, of x] | 
| 369 | obtain d where d: | |
| 370 | "x = fst (normalize_quot (x, y)) * d" | |
| 371 | "y = snd (normalize_quot (x, y)) * d" | |
| 372 | "d dvd x" "d dvd y" "d \<noteq> 0" . | |
| 373 | note d = this[folded assms(2,3)] | |
| 63500 | 374 | from assms have "coprime x' y'" "unit_factor y' = 1" by (simp_all add: coprime_normalize_quot) | 
| 67051 | 375 | with d \<open>is_unit u\<close> show ?thesis | 
| 376 | by (auto simp add: normalized_fracts_def is_unit_left_imp_coprime unit_div_eq_0_iff intro: normalize_quotI) | |
| 63500 | 377 | qed (simp_all add: assms) | 
| 378 | ||
| 379 | lemma normalize_quot_normalize_left: | |
| 380 | fixes x y u | |
| 381 | defines "x' \<equiv> fst (normalize_quot (x, y))" and "y' \<equiv> snd (normalize_quot (x, y))" | |
| 382 | shows "normalize_quot (normalize x, y) = (x' div unit_factor x, y')" | |
| 383 | using normalize_quot_div_unit_left[of "unit_factor x" x y] | |
| 384 | by (cases "x = 0") (simp_all add: assms) | |
| 385 | ||
| 386 | lemma normalize_quot_normalize_right: | |
| 387 | fixes x y u | |
| 388 | defines "x' \<equiv> fst (normalize_quot (x, y))" and "y' \<equiv> snd (normalize_quot (x, y))" | |
| 389 | shows "normalize_quot (x, normalize y) = (x' * unit_factor y, y')" | |
| 390 | using normalize_quot_div_unit_right[of "unit_factor y" x y] | |
| 391 | by (cases "y = 0") (simp_all add: assms) | |
| 392 | ||
| 393 | lemma quot_of_fract_0 [simp]: "quot_of_fract 0 = (0, 1)" | |
| 394 | by transfer auto | |
| 395 | ||
| 396 | lemma quot_of_fract_1 [simp]: "quot_of_fract 1 = (1, 1)" | |
| 397 | by transfer (rule normalize_quotI, simp_all add: normalized_fracts_def) | |
| 398 | ||
| 399 | lemma quot_of_fract_divide: | |
| 400 | "quot_of_fract (x / y) = (if y = 0 then (0, 1) else | |
| 401 | (let (a,b) = quot_of_fract x; (c,d) = quot_of_fract y; | |
| 402 | (e,f) = normalize_quot (a,c); (g,h) = normalize_quot (d,b) | |
| 403 | in (e * g, f * h)))" (is "_ = ?rhs") | |
| 404 | proof (cases "y = 0") | |
| 405 | case False | |
| 406 | hence A: "fst (quot_of_fract y) \<noteq> 0" by transfer auto | |
| 407 | have "x / y = x * inverse y" by (simp add: divide_inverse) | |
| 408 | also from False A have "quot_of_fract \<dots> = ?rhs" | |
| 409 | by (simp only: quot_of_fract_mult quot_of_fract_inverse) | |
| 410 | (simp_all add: Let_def case_prod_unfold fst_quot_of_fract_0_imp | |
| 411 | normalize_quot_div_unit_left normalize_quot_div_unit_right | |
| 412 | normalize_quot_normalize_right normalize_quot_normalize_left) | |
| 413 | finally show ?thesis . | |
| 414 | qed simp_all | |
| 415 | ||
| 416 | end |