| author | wenzelm | 
| Wed, 23 Mar 2022 12:21:13 +0100 | |
| changeset 75311 | 5960bae73afe | 
| parent 69597 | ff784d5a5bfb | 
| child 80105 | 2fa018321400 | 
| permissions | -rw-r--r-- | 
| 54220 
0e6645622f22
more convenient place for a theory in solitariness
 haftmann parents: 
50282diff
changeset | 1 | (* Title: HOL/Decision_Procs/Rat_Pair.thy | 
| 24197 | 2 | Author: Amine Chaieb | 
| 3 | *) | |
| 4 | ||
| 60533 | 5 | section \<open>Rational numbers as pairs\<close> | 
| 24197 | 6 | |
| 54220 
0e6645622f22
more convenient place for a theory in solitariness
 haftmann parents: 
50282diff
changeset | 7 | theory Rat_Pair | 
| 67123 | 8 | imports Complex_Main | 
| 24197 | 9 | begin | 
| 10 | ||
| 42463 | 11 | type_synonym Num = "int \<times> int" | 
| 25005 | 12 | |
| 69597 | 13 | abbreviation Num0_syn :: Num (\<open>0\<^sub>N\<close>) | 
| 44779 | 14 | where "0\<^sub>N \<equiv> (0, 0)" | 
| 25005 | 15 | |
| 69597 | 16 | abbreviation Numi_syn :: "int \<Rightarrow> Num" (\<open>'((_)')\<^sub>N\<close>) | 
| 50282 
fe4d4bb9f4c2
more robust syntax that survives collapse of \<^isub> and \<^sub>;
 wenzelm parents: 
47162diff
changeset | 17 | where "(i)\<^sub>N \<equiv> (i, 1)" | 
| 24197 | 18 | |
| 60538 | 19 | definition isnormNum :: "Num \<Rightarrow> bool" | 
| 20 | where "isnormNum = (\<lambda>(a, b). if a = 0 then b = 0 else b > 0 \<and> gcd a b = 1)" | |
| 24197 | 21 | |
| 60538 | 22 | definition normNum :: "Num \<Rightarrow> Num" | 
| 67123 | 23 | where "normNum = (\<lambda>(a,b). | 
| 60538 | 24 | (if a = 0 \<or> b = 0 then (0, 0) | 
| 25 | else | |
| 44780 | 26 | (let g = gcd a b | 
| 44779 | 27 | in if b > 0 then (a div g, b div g) else (- (a div g), - (b div g)))))" | 
| 24197 | 28 | |
| 62348 | 29 | declare gcd_dvd1[presburger] gcd_dvd2[presburger] | 
| 44779 | 30 | |
| 24197 | 31 | lemma normNum_isnormNum [simp]: "isnormNum (normNum x)" | 
| 32 | proof - | |
| 44780 | 33 | obtain a b where x: "x = (a, b)" by (cases x) | 
| 60567 | 34 | consider "a = 0 \<or> b = 0" | "a \<noteq> 0" "b \<noteq> 0" | 
| 35 | by blast | |
| 60538 | 36 | then show ?thesis | 
| 37 | proof cases | |
| 38 | case 1 | |
| 39 | then show ?thesis | |
| 40 | by (simp add: x normNum_def isnormNum_def) | |
| 41 | next | |
| 60567 | 42 | case ab: 2 | 
| 31706 | 43 | let ?g = "gcd a b" | 
| 24197 | 44 | let ?a' = "a div ?g" | 
| 45 | let ?b' = "b div ?g" | |
| 31706 | 46 | let ?g' = "gcd ?a' ?b'" | 
| 60567 | 47 | from ab have "?g \<noteq> 0" by simp | 
| 48 | with gcd_ge_0_int[of a b] have gpos: "?g > 0" by arith | |
| 44779 | 49 | have gdvd: "?g dvd a" "?g dvd b" by arith+ | 
| 60567 | 50 | from dvd_mult_div_cancel[OF gdvd(1)] dvd_mult_div_cancel[OF gdvd(2)] ab | 
| 44780 | 51 | have nz': "?a' \<noteq> 0" "?b' \<noteq> 0" by - (rule notI, simp)+ | 
| 60567 | 52 | from ab have stupid: "a \<noteq> 0 \<or> b \<noteq> 0" by arith | 
| 67051 | 53 | from div_gcd_coprime[OF stupid] have gp1: "?g' = 1" | 
| 54 | by simp | |
| 60567 | 55 | from ab consider "b < 0" | "b > 0" by arith | 
| 60538 | 56 | then show ?thesis | 
| 57 | proof cases | |
| 60567 | 58 | case b: 1 | 
| 60538 | 59 | have False if b': "?b' \<ge> 0" | 
| 60 | proof - | |
| 61 | from gpos have th: "?g \<ge> 0" by arith | |
| 62 | from mult_nonneg_nonneg[OF th b'] dvd_mult_div_cancel[OF gdvd(2)] | |
| 60567 | 63 | show ?thesis using b by arith | 
| 60538 | 64 | qed | 
| 65 | then have b': "?b' < 0" by (presburger add: linorder_not_le[symmetric]) | |
| 60567 | 66 | from ab(1) nz' b b' gp1 show ?thesis | 
| 60538 | 67 | by (simp add: x isnormNum_def normNum_def Let_def split_def) | 
| 68 | next | |
| 60567 | 69 | case b: 2 | 
| 60538 | 70 | then have "?b' \<ge> 0" | 
| 44779 | 71 | by (presburger add: pos_imp_zdiv_nonneg_iff[OF gpos]) | 
| 44780 | 72 | with nz' have b': "?b' > 0" by arith | 
| 60567 | 73 | from b b' ab(1) nz' gp1 show ?thesis | 
| 60538 | 74 | by (simp add: x isnormNum_def normNum_def Let_def split_def) | 
| 75 | qed | |
| 76 | qed | |
| 24197 | 77 | qed | 
| 78 | ||
| 60533 | 79 | text \<open>Arithmetic over Num\<close> | 
| 24197 | 80 | |
| 69597 | 81 | definition Nadd :: "Num \<Rightarrow> Num \<Rightarrow> Num" (infixl \<open>+\<^sub>N\<close> 60) | 
| 60538 | 82 | where | 
| 83 | "Nadd = (\<lambda>(a, b) (a', b'). | |
| 84 | if a = 0 \<or> b = 0 then normNum (a', b') | |
| 85 | else if a' = 0 \<or> b' = 0 then normNum (a, b) | |
| 86 | else normNum (a * b' + b * a', b * b'))" | |
| 24197 | 87 | |
| 69597 | 88 | definition Nmul :: "Num \<Rightarrow> Num \<Rightarrow> Num" (infixl \<open>*\<^sub>N\<close> 60) | 
| 60538 | 89 | where | 
| 90 | "Nmul = (\<lambda>(a, b) (a', b'). | |
| 91 | let g = gcd (a * a') (b * b') | |
| 92 | in (a * a' div g, b * b' div g))" | |
| 24197 | 93 | |
| 69597 | 94 | definition Nneg :: "Num \<Rightarrow> Num" (\<open>~\<^sub>N\<close>) | 
| 60538 | 95 | where "Nneg = (\<lambda>(a, b). (- a, b))" | 
| 24197 | 96 | |
| 69597 | 97 | definition Nsub :: "Num \<Rightarrow> Num \<Rightarrow> Num" (infixl \<open>-\<^sub>N\<close> 60) | 
| 44779 | 98 | where "Nsub = (\<lambda>a b. a +\<^sub>N ~\<^sub>N b)" | 
| 24197 | 99 | |
| 44779 | 100 | definition Ninv :: "Num \<Rightarrow> Num" | 
| 60538 | 101 | where "Ninv = (\<lambda>(a, b). if a < 0 then (- b, \<bar>a\<bar>) else (b, a))" | 
| 24197 | 102 | |
| 69597 | 103 | definition Ndiv :: "Num \<Rightarrow> Num \<Rightarrow> Num" (infixl \<open>\<div>\<^sub>N\<close> 60) | 
| 44779 | 104 | where "Ndiv = (\<lambda>a b. a *\<^sub>N Ninv b)" | 
| 24197 | 105 | |
| 106 | lemma Nneg_normN[simp]: "isnormNum x \<Longrightarrow> isnormNum (~\<^sub>N x)" | |
| 44779 | 107 | by (simp add: isnormNum_def Nneg_def split_def) | 
| 108 | ||
| 24197 | 109 | lemma Nadd_normN[simp]: "isnormNum (x +\<^sub>N y)" | 
| 110 | by (simp add: Nadd_def split_def) | |
| 44779 | 111 | |
| 60538 | 112 | lemma Nsub_normN[simp]: "isnormNum y \<Longrightarrow> isnormNum (x -\<^sub>N y)" | 
| 24197 | 113 | by (simp add: Nsub_def split_def) | 
| 44779 | 114 | |
| 115 | lemma Nmul_normN[simp]: | |
| 60538 | 116 | assumes xn: "isnormNum x" | 
| 117 | and yn: "isnormNum y" | |
| 24197 | 118 | shows "isnormNum (x *\<^sub>N y)" | 
| 44779 | 119 | proof - | 
| 44780 | 120 | obtain a b where x: "x = (a, b)" by (cases x) | 
| 121 | obtain a' b' where y: "y = (a', b')" by (cases y) | |
| 60538 | 122 | consider "a = 0" | "a' = 0" | "a \<noteq> 0" "a' \<noteq> 0" by blast | 
| 123 | then show ?thesis | |
| 124 | proof cases | |
| 125 | case 1 | |
| 126 | then show ?thesis | |
| 127 | using xn x y by (simp add: isnormNum_def Let_def Nmul_def split_def) | |
| 128 | next | |
| 129 | case 2 | |
| 130 | then show ?thesis | |
| 131 | using yn x y by (simp add: isnormNum_def Let_def Nmul_def split_def) | |
| 132 | next | |
| 60567 | 133 | case aa': 3 | 
| 60538 | 134 | then have bp: "b > 0" "b' > 0" | 
| 135 | using xn yn x y by (simp_all add: isnormNum_def) | |
| 56544 | 136 | from bp have "x *\<^sub>N y = normNum (a * a', b * b')" | 
| 60567 | 137 | using x y aa' bp by (simp add: Nmul_def Let_def split_def normNum_def) | 
| 60538 | 138 | then show ?thesis by simp | 
| 139 | qed | |
| 24197 | 140 | qed | 
| 141 | ||
| 142 | lemma Ninv_normN[simp]: "isnormNum x \<Longrightarrow> isnormNum (Ninv x)" | |
| 60538 | 143 | apply (simp add: Ninv_def isnormNum_def split_def) | 
| 144 | apply (cases "fst x = 0") | |
| 62348 | 145 | apply (auto simp add: gcd.commute) | 
| 60538 | 146 | done | 
| 24197 | 147 | |
| 60538 | 148 | lemma isnormNum_int[simp]: "isnormNum 0\<^sub>N" "isnormNum ((1::int)\<^sub>N)" "i \<noteq> 0 \<Longrightarrow> isnormNum (i)\<^sub>N" | 
| 31706 | 149 | by (simp_all add: isnormNum_def) | 
| 24197 | 150 | |
| 151 | ||
| 60533 | 152 | text \<open>Relations over Num\<close> | 
| 24197 | 153 | |
| 69597 | 154 | definition Nlt0:: "Num \<Rightarrow> bool" (\<open>0>\<^sub>N\<close>) | 
| 60538 | 155 | where "Nlt0 = (\<lambda>(a, b). a < 0)" | 
| 24197 | 156 | |
| 69597 | 157 | definition Nle0:: "Num \<Rightarrow> bool" (\<open>0\<ge>\<^sub>N\<close>) | 
| 60538 | 158 | where "Nle0 = (\<lambda>(a, b). a \<le> 0)" | 
| 24197 | 159 | |
| 69597 | 160 | definition Ngt0:: "Num \<Rightarrow> bool" (\<open>0<\<^sub>N\<close>) | 
| 60538 | 161 | where "Ngt0 = (\<lambda>(a, b). a > 0)" | 
| 24197 | 162 | |
| 69597 | 163 | definition Nge0:: "Num \<Rightarrow> bool" (\<open>0\<le>\<^sub>N\<close>) | 
| 60538 | 164 | where "Nge0 = (\<lambda>(a, b). a \<ge> 0)" | 
| 24197 | 165 | |
| 69597 | 166 | definition Nlt :: "Num \<Rightarrow> Num \<Rightarrow> bool" (infix \<open><\<^sub>N\<close> 55) | 
| 44779 | 167 | where "Nlt = (\<lambda>a b. 0>\<^sub>N (a -\<^sub>N b))" | 
| 24197 | 168 | |
| 69597 | 169 | definition Nle :: "Num \<Rightarrow> Num \<Rightarrow> bool" (infix \<open>\<le>\<^sub>N\<close> 55) | 
| 44779 | 170 | where "Nle = (\<lambda>a b. 0\<ge>\<^sub>N (a -\<^sub>N b))" | 
| 24197 | 171 | |
| 60538 | 172 | definition "INum = (\<lambda>(a, b). of_int a / of_int b)" | 
| 24197 | 173 | |
| 60538 | 174 | lemma INum_int [simp]: "INum (i)\<^sub>N = (of_int i ::'a::field)" "INum 0\<^sub>N = (0::'a::field)" | 
| 24197 | 175 | by (simp_all add: INum_def) | 
| 176 | ||
| 44780 | 177 | lemma isnormNum_unique[simp]: | 
| 60538 | 178 | assumes na: "isnormNum x" | 
| 179 | and nb: "isnormNum y" | |
| 68442 | 180 | shows "(INum x ::'a::field_char_0) = INum y \<longleftrightarrow> x = y" | 
| 60538 | 181 | (is "?lhs = ?rhs") | 
| 24197 | 182 | proof | 
| 44780 | 183 | obtain a b where x: "x = (a, b)" by (cases x) | 
| 184 | obtain a' b' where y: "y = (a', b')" by (cases y) | |
| 60538 | 185 | consider "a = 0 \<or> b = 0 \<or> a' = 0 \<or> b' = 0" | "a \<noteq> 0" "b \<noteq> 0" "a' \<noteq> 0" "b' \<noteq> 0" | 
| 186 | by blast | |
| 187 | then show ?rhs if H: ?lhs | |
| 188 | proof cases | |
| 189 | case 1 | |
| 190 | then show ?thesis | |
| 62390 | 191 | using na nb H by (simp add: x y INum_def split_def isnormNum_def split: if_split_asm) | 
| 60538 | 192 | next | 
| 193 | case 2 | |
| 194 | with na nb have pos: "b > 0" "b' > 0" | |
| 195 | by (simp_all add: x y isnormNum_def) | |
| 60567 | 196 | from H \<open>b \<noteq> 0\<close> \<open>b' \<noteq> 0\<close> have eq: "a * b' = a' * b" | 
| 44780 | 197 | by (simp add: x y INum_def eq_divide_eq divide_eq_eq of_int_mult[symmetric] del: of_int_mult) | 
| 60538 | 198 | from \<open>a \<noteq> 0\<close> \<open>a' \<noteq> 0\<close> na nb | 
| 199 | have gcd1: "gcd a b = 1" "gcd b a = 1" "gcd a' b' = 1" "gcd b' a' = 1" | |
| 62348 | 200 | by (simp_all add: x y isnormNum_def add: gcd.commute) | 
| 67051 | 201 | then have "coprime a b" "coprime b a" "coprime a' b'" "coprime b' a'" | 
| 202 | by (simp_all add: coprime_iff_gcd_eq_1) | |
| 44780 | 203 | from eq have raw_dvd: "a dvd a' * b" "b dvd b' * a" "a' dvd a * b'" "b' dvd b * a'" | 
| 204 | apply - | |
| 27668 | 205 | apply algebra | 
| 206 | apply algebra | |
| 207 | apply simp | |
| 208 | apply algebra | |
| 24197 | 209 | done | 
| 67051 | 210 | then have eq1: "b = b'" | 
| 211 | using pos \<open>coprime b a\<close> \<open>coprime b' a'\<close> | |
| 212 | by (auto simp add: coprime_dvd_mult_left_iff intro: associated_eqI) | |
| 213 | with eq have "a = a'" using pos by simp | |
| 214 | with \<open>b = b'\<close> show ?thesis by (simp add: x y) | |
| 60538 | 215 | qed | 
| 216 | show ?lhs if ?rhs | |
| 217 | using that by simp | |
| 24197 | 218 | qed | 
| 219 | ||
| 68442 | 220 | lemma isnormNum0[simp]: "isnormNum x \<Longrightarrow> INum x = (0::'a::field_char_0) \<longleftrightarrow> x = 0\<^sub>N" | 
| 24197 | 221 | unfolding INum_int(2)[symmetric] | 
| 44779 | 222 | by (rule isnormNum_unique) simp_all | 
| 24197 | 223 | |
| 60538 | 224 | lemma of_int_div_aux: | 
| 225 | assumes "d \<noteq> 0" | |
| 226 | shows "(of_int x ::'a::field_char_0) / of_int d = | |
| 227 | of_int (x div d) + (of_int (x mod d)) / of_int d" | |
| 24197 | 228 | proof - | 
| 60538 | 229 | let ?t = "of_int (x div d) * (of_int d ::'a) + of_int (x mod d)" | 
| 24197 | 230 | let ?f = "\<lambda>x. x / of_int d" | 
| 231 | have "x = (x div d) * d + x mod d" | |
| 232 | by auto | |
| 233 | then have eq: "of_int x = ?t" | |
| 234 | by (simp only: of_int_mult[symmetric] of_int_add [symmetric]) | |
| 44780 | 235 | then have "of_int x / of_int d = ?t / of_int d" | 
| 24197 | 236 | using cong[OF refl[of ?f] eq] by simp | 
| 60538 | 237 | then show ?thesis | 
| 238 | by (simp add: add_divide_distrib algebra_simps \<open>d \<noteq> 0\<close>) | |
| 24197 | 239 | qed | 
| 240 | ||
| 60538 | 241 | lemma of_int_div: | 
| 242 | fixes d :: int | |
| 243 | assumes "d \<noteq> 0" "d dvd n" | |
| 244 | shows "(of_int (n div d) ::'a::field_char_0) = of_int n / of_int d" | |
| 245 | using assms of_int_div_aux [of d n, where ?'a = 'a] by simp | |
| 24197 | 246 | |
| 68442 | 247 | lemma normNum[simp]: "INum (normNum x) = (INum x :: 'a::field_char_0)" | 
| 44779 | 248 | proof - | 
| 44780 | 249 | obtain a b where x: "x = (a, b)" by (cases x) | 
| 60538 | 250 | consider "a = 0 \<or> b = 0" | "a \<noteq> 0" "b \<noteq> 0" by blast | 
| 251 | then show ?thesis | |
| 252 | proof cases | |
| 253 | case 1 | |
| 254 | then show ?thesis | |
| 255 | by (simp add: x INum_def normNum_def split_def Let_def) | |
| 256 | next | |
| 60567 | 257 | case ab: 2 | 
| 31706 | 258 | let ?g = "gcd a b" | 
| 60567 | 259 | from ab have g: "?g \<noteq> 0"by simp | 
| 60538 | 260 | from of_int_div[OF g, where ?'a = 'a] show ?thesis | 
| 261 | by (auto simp add: x INum_def normNum_def split_def Let_def) | |
| 262 | qed | |
| 24197 | 263 | qed | 
| 264 | ||
| 68442 | 265 | lemma INum_normNum_iff: "(INum x ::'a::field_char_0) = INum y \<longleftrightarrow> normNum x = normNum y" | 
| 60538 | 266 | (is "?lhs \<longleftrightarrow> _") | 
| 24197 | 267 | proof - | 
| 268 | have "normNum x = normNum y \<longleftrightarrow> (INum (normNum x) :: 'a) = INum (normNum y)" | |
| 269 | by (simp del: normNum) | |
| 270 | also have "\<dots> = ?lhs" by simp | |
| 271 | finally show ?thesis by simp | |
| 272 | qed | |
| 273 | ||
| 68442 | 274 | lemma Nadd[simp]: "INum (x +\<^sub>N y) = INum x + (INum y :: 'a :: field_char_0)" | 
| 44779 | 275 | proof - | 
| 24197 | 276 | let ?z = "0::'a" | 
| 44780 | 277 | obtain a b where x: "x = (a, b)" by (cases x) | 
| 278 | obtain a' b' where y: "y = (a', b')" by (cases y) | |
| 60538 | 279 | consider "a = 0 \<or> a'= 0 \<or> b =0 \<or> b' = 0" | "a \<noteq> 0" "a'\<noteq> 0" "b \<noteq> 0" "b' \<noteq> 0" | 
| 280 | by blast | |
| 281 | then show ?thesis | |
| 282 | proof cases | |
| 283 | case 1 | |
| 284 | then show ?thesis | |
| 285 | apply (cases "a = 0") | |
| 286 | apply (simp_all add: x y Nadd_def) | |
| 287 | apply (cases "b = 0") | |
| 288 | apply (simp_all add: INum_def) | |
| 289 | apply (cases "a'= 0") | |
| 290 | apply simp_all | |
| 291 | apply (cases "b'= 0") | |
| 292 | apply simp_all | |
| 293 | done | |
| 294 | next | |
| 60567 | 295 | case neq: 2 | 
| 60538 | 296 | show ?thesis | 
| 297 | proof (cases "a * b' + b * a' = 0") | |
| 298 | case True | |
| 299 | then have "of_int (a * b' + b * a') / (of_int b * of_int b') = ?z" | |
| 300 | by simp | |
| 301 | then have "of_int b' * of_int a / (of_int b * of_int b') + | |
| 302 | of_int b * of_int a' / (of_int b * of_int b') = ?z" | |
| 303 | by (simp add: add_divide_distrib) | |
| 304 | then have th: "of_int a / of_int b + of_int a' / of_int b' = ?z" | |
| 60567 | 305 | using neq by simp | 
| 306 | from True neq show ?thesis | |
| 60538 | 307 | by (simp add: x y th Nadd_def normNum_def INum_def split_def) | 
| 308 | next | |
| 309 | case False | |
| 310 | let ?g = "gcd (a * b' + b * a') (b * b')" | |
| 311 | have gz: "?g \<noteq> 0" | |
| 312 | using False by simp | |
| 313 | show ?thesis | |
| 60567 | 314 | using neq False gz | 
| 60538 | 315 | of_int_div [where ?'a = 'a, OF gz gcd_dvd1 [of "a * b' + b * a'" "b * b'"]] | 
| 316 | of_int_div [where ?'a = 'a, OF gz gcd_dvd2 [of "a * b' + b * a'" "b * b'"]] | |
| 317 | by (simp add: x y Nadd_def INum_def normNum_def Let_def) (simp add: field_simps) | |
| 318 | qed | |
| 319 | qed | |
| 24197 | 320 | qed | 
| 321 | ||
| 68442 | 322 | lemma Nmul[simp]: "INum (x *\<^sub>N y) = INum x * (INum y:: 'a::field_char_0)" | 
| 60538 | 323 | proof - | 
| 324 | let ?z = "0::'a" | |
| 325 | obtain a b where x: "x = (a, b)" by (cases x) | |
| 326 | obtain a' b' where y: "y = (a', b')" by (cases y) | |
| 327 | consider "a = 0 \<or> a' = 0 \<or> b = 0 \<or> b' = 0" | "a \<noteq> 0" "a' \<noteq> 0" "b \<noteq> 0" "b' \<noteq> 0" | |
| 328 | by blast | |
| 329 | then show ?thesis | |
| 330 | proof cases | |
| 331 | case 1 | |
| 332 | then show ?thesis | |
| 60698 | 333 | by (auto simp add: x y Nmul_def INum_def) | 
| 60538 | 334 | next | 
| 60567 | 335 | case neq: 2 | 
| 60538 | 336 | let ?g = "gcd (a * a') (b * b')" | 
| 337 | have gz: "?g \<noteq> 0" | |
| 60567 | 338 | using neq by simp | 
| 339 | from neq of_int_div [where ?'a = 'a, OF gz gcd_dvd1 [of "a * a'" "b * b'"]] | |
| 60538 | 340 | of_int_div [where ?'a = 'a , OF gz gcd_dvd2 [of "a * a'" "b * b'"]] | 
| 341 | show ?thesis | |
| 342 | by (simp add: Nmul_def x y Let_def INum_def) | |
| 343 | qed | |
| 344 | qed | |
| 345 | ||
| 60698 | 346 | lemma Nneg[simp]: "INum (~\<^sub>N x) = - (INum x :: 'a::field)" | 
| 56479 
91958d4b30f7
revert c1bbd3e22226, a14831ac3023, and 36489d77c484: divide_minus_left/right are again simp rules
 hoelzl parents: 
56410diff
changeset | 347 | by (simp add: Nneg_def split_def INum_def) | 
| 24197 | 348 | |
| 68442 | 349 | lemma Nsub[simp]: "INum (x -\<^sub>N y) = INum x - (INum y:: 'a::field_char_0)" | 
| 44779 | 350 | by (simp add: Nsub_def split_def) | 
| 24197 | 351 | |
| 60698 | 352 | lemma Ninv[simp]: "INum (Ninv x) = (1 :: 'a::field) / INum x" | 
| 56479 
91958d4b30f7
revert c1bbd3e22226, a14831ac3023, and 36489d77c484: divide_minus_left/right are again simp rules
 hoelzl parents: 
56410diff
changeset | 353 | by (simp add: Ninv_def INum_def split_def) | 
| 24197 | 354 | |
| 68442 | 355 | lemma Ndiv[simp]: "INum (x \<div>\<^sub>N y) = INum x / (INum y :: 'a::field_char_0)" | 
| 44779 | 356 | by (simp add: Ndiv_def) | 
| 24197 | 357 | |
| 44779 | 358 | lemma Nlt0_iff[simp]: | 
| 44780 | 359 | assumes nx: "isnormNum x" | 
| 60698 | 360 |   shows "((INum x :: 'a::{field_char_0,linordered_field}) < 0) = 0>\<^sub>N x"
 | 
| 44779 | 361 | proof - | 
| 44780 | 362 | obtain a b where x: "x = (a, b)" by (cases x) | 
| 60538 | 363 | show ?thesis | 
| 364 | proof (cases "a = 0") | |
| 365 | case True | |
| 366 | then show ?thesis | |
| 367 | by (simp add: x Nlt0_def INum_def) | |
| 368 | next | |
| 369 | case False | |
| 370 | then have b: "(of_int b::'a) > 0" | |
| 44780 | 371 | using nx by (simp add: x isnormNum_def) | 
| 24197 | 372 | from pos_divide_less_eq[OF b, where b="of_int a" and a="0::'a"] | 
| 60538 | 373 | show ?thesis | 
| 374 | by (simp add: x Nlt0_def INum_def) | |
| 375 | qed | |
| 24197 | 376 | qed | 
| 377 | ||
| 44779 | 378 | lemma Nle0_iff[simp]: | 
| 379 | assumes nx: "isnormNum x" | |
| 60538 | 380 |   shows "((INum x :: 'a::{field_char_0,linordered_field}) \<le> 0) = 0\<ge>\<^sub>N x"
 | 
| 44779 | 381 | proof - | 
| 44780 | 382 | obtain a b where x: "x = (a, b)" by (cases x) | 
| 60538 | 383 | show ?thesis | 
| 384 | proof (cases "a = 0") | |
| 385 | case True | |
| 386 | then show ?thesis | |
| 387 | by (simp add: x Nle0_def INum_def) | |
| 388 | next | |
| 389 | case False | |
| 390 | then have b: "(of_int b :: 'a) > 0" | |
| 44780 | 391 | using nx by (simp add: x isnormNum_def) | 
| 24197 | 392 | from pos_divide_le_eq[OF b, where b="of_int a" and a="0::'a"] | 
| 60538 | 393 | show ?thesis | 
| 394 | by (simp add: x Nle0_def INum_def) | |
| 395 | qed | |
| 24197 | 396 | qed | 
| 397 | ||
| 44779 | 398 | lemma Ngt0_iff[simp]: | 
| 399 | assumes nx: "isnormNum x" | |
| 60698 | 400 |   shows "((INum x :: 'a::{field_char_0,linordered_field}) > 0) = 0<\<^sub>N x"
 | 
| 44779 | 401 | proof - | 
| 44780 | 402 | obtain a b where x: "x = (a, b)" by (cases x) | 
| 60538 | 403 | show ?thesis | 
| 404 | proof (cases "a = 0") | |
| 405 | case True | |
| 406 | then show ?thesis | |
| 407 | by (simp add: x Ngt0_def INum_def) | |
| 408 | next | |
| 409 | case False | |
| 410 | then have b: "(of_int b::'a) > 0" | |
| 411 | using nx by (simp add: x isnormNum_def) | |
| 24197 | 412 | from pos_less_divide_eq[OF b, where b="of_int a" and a="0::'a"] | 
| 60538 | 413 | show ?thesis | 
| 414 | by (simp add: x Ngt0_def INum_def) | |
| 415 | qed | |
| 24197 | 416 | qed | 
| 417 | ||
| 44779 | 418 | lemma Nge0_iff[simp]: | 
| 419 | assumes nx: "isnormNum x" | |
| 60698 | 420 |   shows "(INum x :: 'a::{field_char_0,linordered_field}) \<ge> 0 \<longleftrightarrow> 0\<le>\<^sub>N x"
 | 
| 44779 | 421 | proof - | 
| 44780 | 422 | obtain a b where x: "x = (a, b)" by (cases x) | 
| 60538 | 423 | show ?thesis | 
| 424 | proof (cases "a = 0") | |
| 425 | case True | |
| 426 | then show ?thesis | |
| 427 | by (simp add: x Nge0_def INum_def) | |
| 428 | next | |
| 429 | case False | |
| 430 | then have b: "(of_int b::'a) > 0" | |
| 431 | using nx by (simp add: x isnormNum_def) | |
| 44779 | 432 | from pos_le_divide_eq[OF b, where b="of_int a" and a="0::'a"] | 
| 60538 | 433 | show ?thesis | 
| 434 | by (simp add: x Nge0_def INum_def) | |
| 435 | qed | |
| 44779 | 436 | qed | 
| 437 | ||
| 438 | lemma Nlt_iff[simp]: | |
| 60538 | 439 | assumes nx: "isnormNum x" | 
| 440 | and ny: "isnormNum y" | |
| 60698 | 441 |   shows "((INum x :: 'a::{field_char_0,linordered_field}) < INum y) \<longleftrightarrow> x <\<^sub>N y"
 | 
| 44779 | 442 | proof - | 
| 24197 | 443 | let ?z = "0::'a" | 
| 60698 | 444 | have "((INum x ::'a) < INum y) \<longleftrightarrow> INum (x -\<^sub>N y) < ?z" | 
| 44779 | 445 | using nx ny by simp | 
| 60698 | 446 | also have "\<dots> \<longleftrightarrow> 0>\<^sub>N (x -\<^sub>N y)" | 
| 44779 | 447 | using Nlt0_iff[OF Nsub_normN[OF ny]] by simp | 
| 60538 | 448 | finally show ?thesis | 
| 449 | by (simp add: Nlt_def) | |
| 24197 | 450 | qed | 
| 451 | ||
| 44779 | 452 | lemma Nle_iff[simp]: | 
| 60538 | 453 | assumes nx: "isnormNum x" | 
| 454 | and ny: "isnormNum y" | |
| 60698 | 455 |   shows "((INum x :: 'a::{field_char_0,linordered_field}) \<le> INum y) \<longleftrightarrow> x \<le>\<^sub>N y"
 | 
| 44779 | 456 | proof - | 
| 60698 | 457 | have "((INum x ::'a) \<le> INum y) \<longleftrightarrow> INum (x -\<^sub>N y) \<le> (0::'a)" | 
| 44779 | 458 | using nx ny by simp | 
| 60698 | 459 | also have "\<dots> \<longleftrightarrow> 0\<ge>\<^sub>N (x -\<^sub>N y)" | 
| 44779 | 460 | using Nle0_iff[OF Nsub_normN[OF ny]] by simp | 
| 60538 | 461 | finally show ?thesis | 
| 462 | by (simp add: Nle_def) | |
| 24197 | 463 | qed | 
| 464 | ||
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 465 | lemma Nadd_commute: | 
| 68442 | 466 |   assumes "SORT_CONSTRAINT('a::field_char_0)"
 | 
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 467 | shows "x +\<^sub>N y = y +\<^sub>N x" | 
| 44779 | 468 | proof - | 
| 60538 | 469 | have n: "isnormNum (x +\<^sub>N y)" "isnormNum (y +\<^sub>N x)" | 
| 470 | by simp_all | |
| 471 | have "(INum (x +\<^sub>N y)::'a) = INum (y +\<^sub>N x)" | |
| 472 | by simp | |
| 473 | with isnormNum_unique[OF n] show ?thesis | |
| 474 | by simp | |
| 24197 | 475 | qed | 
| 476 | ||
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 477 | lemma [simp]: | 
| 68442 | 478 |   assumes "SORT_CONSTRAINT('a::field_char_0)"
 | 
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 479 | shows "(0, b) +\<^sub>N y = normNum y" | 
| 44780 | 480 | and "(a, 0) +\<^sub>N y = normNum y" | 
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 481 | and "x +\<^sub>N (0, b) = normNum x" | 
| 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 482 | and "x +\<^sub>N (a, 0) = normNum x" | 
| 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 483 | apply (simp add: Nadd_def split_def) | 
| 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 484 | apply (simp add: Nadd_def split_def) | 
| 60538 | 485 | apply (subst Nadd_commute) | 
| 486 | apply (simp add: Nadd_def split_def) | |
| 487 | apply (subst Nadd_commute) | |
| 488 | apply (simp add: Nadd_def split_def) | |
| 24197 | 489 | done | 
| 490 | ||
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 491 | lemma normNum_nilpotent_aux[simp]: | 
| 68442 | 492 |   assumes "SORT_CONSTRAINT('a::field_char_0)"
 | 
| 44780 | 493 | assumes nx: "isnormNum x" | 
| 24197 | 494 | shows "normNum x = x" | 
| 44779 | 495 | proof - | 
| 24197 | 496 | let ?a = "normNum x" | 
| 497 | have n: "isnormNum ?a" by simp | |
| 44779 | 498 | have th: "INum ?a = (INum x ::'a)" by simp | 
| 499 | with isnormNum_unique[OF n nx] show ?thesis by simp | |
| 24197 | 500 | qed | 
| 501 | ||
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 502 | lemma normNum_nilpotent[simp]: | 
| 68442 | 503 |   assumes "SORT_CONSTRAINT('a::field_char_0)"
 | 
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 504 | shows "normNum (normNum x) = normNum x" | 
| 24197 | 505 | by simp | 
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 506 | |
| 60698 | 507 | lemma normNum0[simp]: "normNum (0, b) = 0\<^sub>N" "normNum (a, 0) = 0\<^sub>N" | 
| 24197 | 508 | by (simp_all add: normNum_def) | 
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 509 | |
| 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 510 | lemma normNum_Nadd: | 
| 68442 | 511 |   assumes "SORT_CONSTRAINT('a::field_char_0)"
 | 
| 60538 | 512 | shows "normNum (x +\<^sub>N y) = x +\<^sub>N y" | 
| 513 | by simp | |
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 514 | |
| 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 515 | lemma Nadd_normNum1[simp]: | 
| 68442 | 516 |   assumes "SORT_CONSTRAINT('a::field_char_0)"
 | 
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 517 | shows "normNum x +\<^sub>N y = x +\<^sub>N y" | 
| 44779 | 518 | proof - | 
| 60698 | 519 | have n: "isnormNum (normNum x +\<^sub>N y)" "isnormNum (x +\<^sub>N y)" | 
| 520 | by simp_all | |
| 521 | have "INum (normNum x +\<^sub>N y) = INum x + (INum y :: 'a)" | |
| 522 | by simp | |
| 523 | also have "\<dots> = INum (x +\<^sub>N y)" | |
| 524 | by simp | |
| 525 | finally show ?thesis | |
| 526 | using isnormNum_unique[OF n] by simp | |
| 24197 | 527 | qed | 
| 528 | ||
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 529 | lemma Nadd_normNum2[simp]: | 
| 68442 | 530 |   assumes "SORT_CONSTRAINT('a::field_char_0)"
 | 
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 531 | shows "x +\<^sub>N normNum y = x +\<^sub>N y" | 
| 44779 | 532 | proof - | 
| 60698 | 533 | have n: "isnormNum (x +\<^sub>N normNum y)" "isnormNum (x +\<^sub>N y)" | 
| 534 | by simp_all | |
| 535 | have "INum (x +\<^sub>N normNum y) = INum x + (INum y :: 'a)" | |
| 536 | by simp | |
| 537 | also have "\<dots> = INum (x +\<^sub>N y)" | |
| 538 | by simp | |
| 539 | finally show ?thesis | |
| 540 | using isnormNum_unique[OF n] by simp | |
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 541 | qed | 
| 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 542 | |
| 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 543 | lemma Nadd_assoc: | 
| 68442 | 544 |   assumes "SORT_CONSTRAINT('a::field_char_0)"
 | 
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 545 | shows "x +\<^sub>N y +\<^sub>N z = x +\<^sub>N (y +\<^sub>N z)" | 
| 44779 | 546 | proof - | 
| 60698 | 547 | have n: "isnormNum (x +\<^sub>N y +\<^sub>N z)" "isnormNum (x +\<^sub>N (y +\<^sub>N z))" | 
| 548 | by simp_all | |
| 549 | have "INum (x +\<^sub>N y +\<^sub>N z) = (INum (x +\<^sub>N (y +\<^sub>N z)) :: 'a)" | |
| 550 | by simp | |
| 551 | with isnormNum_unique[OF n] show ?thesis | |
| 552 | by simp | |
| 24197 | 553 | qed | 
| 554 | ||
| 555 | lemma Nmul_commute: "isnormNum x \<Longrightarrow> isnormNum y \<Longrightarrow> x *\<^sub>N y = y *\<^sub>N x" | |
| 62348 | 556 | by (simp add: Nmul_def split_def Let_def gcd.commute mult.commute) | 
| 24197 | 557 | |
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 558 | lemma Nmul_assoc: | 
| 68442 | 559 |   assumes "SORT_CONSTRAINT('a::field_char_0)"
 | 
| 60538 | 560 | assumes nx: "isnormNum x" | 
| 561 | and ny: "isnormNum y" | |
| 562 | and nz: "isnormNum z" | |
| 24197 | 563 | shows "x *\<^sub>N y *\<^sub>N z = x *\<^sub>N (y *\<^sub>N z)" | 
| 44779 | 564 | proof - | 
| 44780 | 565 | from nx ny nz have n: "isnormNum (x *\<^sub>N y *\<^sub>N z)" "isnormNum (x *\<^sub>N (y *\<^sub>N z))" | 
| 24197 | 566 | by simp_all | 
| 60698 | 567 | have "INum (x +\<^sub>N y +\<^sub>N z) = (INum (x +\<^sub>N (y +\<^sub>N z)) :: 'a)" | 
| 568 | by simp | |
| 569 | with isnormNum_unique[OF n] show ?thesis | |
| 570 | by simp | |
| 24197 | 571 | qed | 
| 572 | ||
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 573 | lemma Nsub0: | 
| 68442 | 574 |   assumes "SORT_CONSTRAINT('a::field_char_0)"
 | 
| 60538 | 575 | assumes x: "isnormNum x" | 
| 576 | and y: "isnormNum y" | |
| 44780 | 577 | shows "x -\<^sub>N y = 0\<^sub>N \<longleftrightarrow> x = y" | 
| 44779 | 578 | proof - | 
| 44780 | 579 | from isnormNum_unique[where 'a = 'a, OF Nsub_normN[OF y], where y="0\<^sub>N"] | 
| 60698 | 580 | have "x -\<^sub>N y = 0\<^sub>N \<longleftrightarrow> INum (x -\<^sub>N y) = (INum 0\<^sub>N :: 'a)" | 
| 581 | by simp | |
| 582 | also have "\<dots> \<longleftrightarrow> INum x = (INum y :: 'a)" | |
| 583 | by simp | |
| 584 | also have "\<dots> \<longleftrightarrow> x = y" | |
| 585 | using x y by simp | |
| 44779 | 586 | finally show ?thesis . | 
| 24197 | 587 | qed | 
| 588 | ||
| 589 | lemma Nmul0[simp]: "c *\<^sub>N 0\<^sub>N = 0\<^sub>N" " 0\<^sub>N *\<^sub>N c = 0\<^sub>N" | |
| 590 | by (simp_all add: Nmul_def Let_def split_def) | |
| 591 | ||
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 592 | lemma Nmul_eq0[simp]: | 
| 68442 | 593 |   assumes "SORT_CONSTRAINT('a::field_char_0)"
 | 
| 60538 | 594 | assumes nx: "isnormNum x" | 
| 595 | and ny: "isnormNum y" | |
| 44780 | 596 | shows "x*\<^sub>N y = 0\<^sub>N \<longleftrightarrow> x = 0\<^sub>N \<or> y = 0\<^sub>N" | 
| 44779 | 597 | proof - | 
| 44780 | 598 | obtain a b where x: "x = (a, b)" by (cases x) | 
| 599 | obtain a' b' where y: "y = (a', b')" by (cases y) | |
| 44779 | 600 | have n0: "isnormNum 0\<^sub>N" by simp | 
| 44780 | 601 | show ?thesis using nx ny | 
| 44779 | 602 | apply (simp only: isnormNum_unique[where ?'a = 'a, OF Nmul_normN[OF nx ny] n0, symmetric] | 
| 603 | Nmul[where ?'a = 'a]) | |
| 62390 | 604 | apply (simp add: x y INum_def split_def isnormNum_def split: if_split_asm) | 
| 44779 | 605 | done | 
| 24197 | 606 | qed | 
| 44779 | 607 | |
| 24197 | 608 | lemma Nneg_Nneg[simp]: "~\<^sub>N (~\<^sub>N c) = c" | 
| 609 | by (simp add: Nneg_def split_def) | |
| 610 | ||
| 60538 | 611 | lemma Nmul1[simp]: "isnormNum c \<Longrightarrow> (1)\<^sub>N *\<^sub>N c = c" "isnormNum c \<Longrightarrow> c *\<^sub>N (1)\<^sub>N = c" | 
| 24197 | 612 | apply (simp_all add: Nmul_def Let_def split_def isnormNum_def) | 
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 613 | apply (cases "fst c = 0", simp_all, cases c, simp_all)+ | 
| 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 614 | done | 
| 24197 | 615 | |
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 616 | end |