| author | wenzelm | 
| Fri, 25 Sep 2015 19:28:33 +0200 | |
| changeset 61265 | 996d73a96b4f | 
| parent 60698 | 29e8bdc41f90 | 
| child 62348 | 9a5f43dac883 | 
| 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 | 
| 36411 | 8 | imports Complex_Main | 
| 24197 | 9 | begin | 
| 10 | ||
| 42463 | 11 | type_synonym Num = "int \<times> int" | 
| 25005 | 12 | |
| 44780 | 13 | abbreviation Num0_syn :: Num  ("0\<^sub>N")
 | 
| 44779 | 14 | where "0\<^sub>N \<equiv> (0, 0)" | 
| 25005 | 15 | |
| 50282 
fe4d4bb9f4c2
more robust syntax that survives collapse of \<^isub> and \<^sub>;
 wenzelm parents: 
47162diff
changeset | 16 | abbreviation Numi_syn :: "int \<Rightarrow> Num"  ("'((_)')\<^sub>N")
 | 
| 
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" | 
| 23 | where | |
| 44779 | 24 | "normNum = (\<lambda>(a,b). | 
| 60538 | 25 | (if a = 0 \<or> b = 0 then (0, 0) | 
| 26 | else | |
| 44780 | 27 | (let g = gcd a b | 
| 44779 | 28 | in if b > 0 then (a div g, b div g) else (- (a div g), - (b div g)))))" | 
| 24197 | 29 | |
| 44779 | 30 | declare gcd_dvd1_int[presburger] gcd_dvd2_int[presburger] | 
| 31 | ||
| 24197 | 32 | lemma normNum_isnormNum [simp]: "isnormNum (normNum x)" | 
| 33 | proof - | |
| 44780 | 34 | obtain a b where x: "x = (a, b)" by (cases x) | 
| 60567 | 35 | consider "a = 0 \<or> b = 0" | "a \<noteq> 0" "b \<noteq> 0" | 
| 36 | by blast | |
| 60538 | 37 | then show ?thesis | 
| 38 | proof cases | |
| 39 | case 1 | |
| 40 | then show ?thesis | |
| 41 | by (simp add: x normNum_def isnormNum_def) | |
| 42 | next | |
| 60567 | 43 | case ab: 2 | 
| 31706 | 44 | let ?g = "gcd a b" | 
| 24197 | 45 | let ?a' = "a div ?g" | 
| 46 | let ?b' = "b div ?g" | |
| 31706 | 47 | let ?g' = "gcd ?a' ?b'" | 
| 60567 | 48 | from ab have "?g \<noteq> 0" by simp | 
| 49 | with gcd_ge_0_int[of a b] have gpos: "?g > 0" by arith | |
| 44779 | 50 | have gdvd: "?g dvd a" "?g dvd b" by arith+ | 
| 60567 | 51 | from dvd_mult_div_cancel[OF gdvd(1)] dvd_mult_div_cancel[OF gdvd(2)] ab | 
| 44780 | 52 | have nz': "?a' \<noteq> 0" "?b' \<noteq> 0" by - (rule notI, simp)+ | 
| 60567 | 53 | from ab have stupid: "a \<noteq> 0 \<or> b \<noteq> 0" by arith | 
| 31952 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 nipkow parents: 
31706diff
changeset | 54 | from div_gcd_coprime_int[OF stupid] have gp1: "?g' = 1" . | 
| 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 | |
| 60538 | 81 | definition Nadd :: "Num \<Rightarrow> Num \<Rightarrow> Num" (infixl "+\<^sub>N" 60) | 
| 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 | |
| 60538 | 88 | definition Nmul :: "Num \<Rightarrow> Num \<Rightarrow> Num" (infixl "*\<^sub>N" 60) | 
| 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 | |
| 44779 | 94 | definition Nneg :: "Num \<Rightarrow> Num" ("~\<^sub>N")
 | 
| 60538 | 95 | where "Nneg = (\<lambda>(a, b). (- a, b))" | 
| 24197 | 96 | |
| 44780 | 97 | definition Nsub :: "Num \<Rightarrow> Num \<Rightarrow> Num" (infixl "-\<^sub>N" 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 | |
| 44780 | 103 | definition Ndiv :: "Num \<Rightarrow> Num \<Rightarrow> Num" (infixl "\<div>\<^sub>N" 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") | |
| 145 | apply (auto simp add: gcd_commute_int) | |
| 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 | |
| 44780 | 154 | definition Nlt0:: "Num \<Rightarrow> bool"  ("0>\<^sub>N")
 | 
| 60538 | 155 | where "Nlt0 = (\<lambda>(a, b). a < 0)" | 
| 24197 | 156 | |
| 44780 | 157 | definition Nle0:: "Num \<Rightarrow> bool"  ("0\<ge>\<^sub>N")
 | 
| 60538 | 158 | where "Nle0 = (\<lambda>(a, b). a \<le> 0)" | 
| 24197 | 159 | |
| 44780 | 160 | definition Ngt0:: "Num \<Rightarrow> bool"  ("0<\<^sub>N")
 | 
| 60538 | 161 | where "Ngt0 = (\<lambda>(a, b). a > 0)" | 
| 24197 | 162 | |
| 44780 | 163 | definition Nge0:: "Num \<Rightarrow> bool"  ("0\<le>\<^sub>N")
 | 
| 60538 | 164 | where "Nge0 = (\<lambda>(a, b). a \<ge> 0)" | 
| 24197 | 165 | |
| 44780 | 166 | definition Nlt :: "Num \<Rightarrow> Num \<Rightarrow> bool" (infix "<\<^sub>N" 55) | 
| 44779 | 167 | where "Nlt = (\<lambda>a b. 0>\<^sub>N (a -\<^sub>N b))" | 
| 24197 | 168 | |
| 44779 | 169 | definition Nle :: "Num \<Rightarrow> Num \<Rightarrow> bool" (infix "\<le>\<^sub>N" 55) | 
| 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" | |
| 180 |   shows "(INum x ::'a::{field_char_0,field}) = INum y \<longleftrightarrow> x = y"
 | |
| 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 | |
| 191 | using na nb H by (simp add: x y INum_def split_def isnormNum_def split: split_if_asm) | |
| 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" | |
| 44780 | 200 | by (simp_all add: x y isnormNum_def add: gcd_commute_int) | 
| 201 | from eq have raw_dvd: "a dvd a' * b" "b dvd b' * a" "a' dvd a * b'" "b' dvd b * a'" | |
| 202 | apply - | |
| 27668 | 203 | apply algebra | 
| 204 | apply algebra | |
| 205 | apply simp | |
| 206 | apply algebra | |
| 24197 | 207 | done | 
| 33657 | 208 | from zdvd_antisym_abs[OF coprime_dvd_mult_int[OF gcd1(2) raw_dvd(2)] | 
| 44780 | 209 | coprime_dvd_mult_int[OF gcd1(4) raw_dvd(4)]] | 
| 41528 | 210 | have eq1: "b = b'" using pos by arith | 
| 24197 | 211 | with eq have "a = a'" using pos by simp | 
| 60538 | 212 | with eq1 show ?thesis by (simp add: x y) | 
| 213 | qed | |
| 214 | show ?lhs if ?rhs | |
| 215 | using that by simp | |
| 24197 | 216 | qed | 
| 217 | ||
| 60538 | 218 | lemma isnormNum0[simp]: "isnormNum x \<Longrightarrow> INum x = (0::'a::{field_char_0,field}) \<longleftrightarrow> x = 0\<^sub>N"
 | 
| 24197 | 219 | unfolding INum_int(2)[symmetric] | 
| 44779 | 220 | by (rule isnormNum_unique) simp_all | 
| 24197 | 221 | |
| 60538 | 222 | lemma of_int_div_aux: | 
| 223 | assumes "d \<noteq> 0" | |
| 224 | shows "(of_int x ::'a::field_char_0) / of_int d = | |
| 225 | of_int (x div d) + (of_int (x mod d)) / of_int d" | |
| 24197 | 226 | proof - | 
| 60538 | 227 | let ?t = "of_int (x div d) * (of_int d ::'a) + of_int (x mod d)" | 
| 24197 | 228 | let ?f = "\<lambda>x. x / of_int d" | 
| 229 | have "x = (x div d) * d + x mod d" | |
| 230 | by auto | |
| 231 | then have eq: "of_int x = ?t" | |
| 232 | by (simp only: of_int_mult[symmetric] of_int_add [symmetric]) | |
| 44780 | 233 | then have "of_int x / of_int d = ?t / of_int d" | 
| 24197 | 234 | using cong[OF refl[of ?f] eq] by simp | 
| 60538 | 235 | then show ?thesis | 
| 236 | by (simp add: add_divide_distrib algebra_simps \<open>d \<noteq> 0\<close>) | |
| 24197 | 237 | qed | 
| 238 | ||
| 60538 | 239 | lemma of_int_div: | 
| 240 | fixes d :: int | |
| 241 | assumes "d \<noteq> 0" "d dvd n" | |
| 242 | shows "(of_int (n div d) ::'a::field_char_0) = of_int n / of_int d" | |
| 243 | using assms of_int_div_aux [of d n, where ?'a = 'a] by simp | |
| 24197 | 244 | |
| 60538 | 245 | lemma normNum[simp]: "INum (normNum x) = (INum x :: 'a::{field_char_0,field})"
 | 
| 44779 | 246 | proof - | 
| 44780 | 247 | obtain a b where x: "x = (a, b)" by (cases x) | 
| 60538 | 248 | consider "a = 0 \<or> b = 0" | "a \<noteq> 0" "b \<noteq> 0" by blast | 
| 249 | then show ?thesis | |
| 250 | proof cases | |
| 251 | case 1 | |
| 252 | then show ?thesis | |
| 253 | by (simp add: x INum_def normNum_def split_def Let_def) | |
| 254 | next | |
| 60567 | 255 | case ab: 2 | 
| 31706 | 256 | let ?g = "gcd a b" | 
| 60567 | 257 | from ab have g: "?g \<noteq> 0"by simp | 
| 60538 | 258 | from of_int_div[OF g, where ?'a = 'a] show ?thesis | 
| 259 | by (auto simp add: x INum_def normNum_def split_def Let_def) | |
| 260 | qed | |
| 24197 | 261 | qed | 
| 262 | ||
| 60538 | 263 | lemma INum_normNum_iff: "(INum x ::'a::{field_char_0,field}) = INum y \<longleftrightarrow> normNum x = normNum y"
 | 
| 264 | (is "?lhs \<longleftrightarrow> _") | |
| 24197 | 265 | proof - | 
| 266 | have "normNum x = normNum y \<longleftrightarrow> (INum (normNum x) :: 'a) = INum (normNum y)" | |
| 267 | by (simp del: normNum) | |
| 268 | also have "\<dots> = ?lhs" by simp | |
| 269 | finally show ?thesis by simp | |
| 270 | qed | |
| 271 | ||
| 60538 | 272 | lemma Nadd[simp]: "INum (x +\<^sub>N y) = INum x + (INum y :: 'a :: {field_char_0,field})"
 | 
| 44779 | 273 | proof - | 
| 24197 | 274 | let ?z = "0::'a" | 
| 44780 | 275 | obtain a b where x: "x = (a, b)" by (cases x) | 
| 276 | obtain a' b' where y: "y = (a', b')" by (cases y) | |
| 60538 | 277 | consider "a = 0 \<or> a'= 0 \<or> b =0 \<or> b' = 0" | "a \<noteq> 0" "a'\<noteq> 0" "b \<noteq> 0" "b' \<noteq> 0" | 
| 278 | by blast | |
| 279 | then show ?thesis | |
| 280 | proof cases | |
| 281 | case 1 | |
| 282 | then show ?thesis | |
| 283 | apply (cases "a = 0") | |
| 284 | apply (simp_all add: x y Nadd_def) | |
| 285 | apply (cases "b = 0") | |
| 286 | apply (simp_all add: INum_def) | |
| 287 | apply (cases "a'= 0") | |
| 288 | apply simp_all | |
| 289 | apply (cases "b'= 0") | |
| 290 | apply simp_all | |
| 291 | done | |
| 292 | next | |
| 60567 | 293 | case neq: 2 | 
| 60538 | 294 | show ?thesis | 
| 295 | proof (cases "a * b' + b * a' = 0") | |
| 296 | case True | |
| 297 | then have "of_int (a * b' + b * a') / (of_int b * of_int b') = ?z" | |
| 298 | by simp | |
| 299 | then have "of_int b' * of_int a / (of_int b * of_int b') + | |
| 300 | of_int b * of_int a' / (of_int b * of_int b') = ?z" | |
| 301 | by (simp add: add_divide_distrib) | |
| 302 | then have th: "of_int a / of_int b + of_int a' / of_int b' = ?z" | |
| 60567 | 303 | using neq by simp | 
| 304 | from True neq show ?thesis | |
| 60538 | 305 | by (simp add: x y th Nadd_def normNum_def INum_def split_def) | 
| 306 | next | |
| 307 | case False | |
| 308 | let ?g = "gcd (a * b' + b * a') (b * b')" | |
| 309 | have gz: "?g \<noteq> 0" | |
| 310 | using False by simp | |
| 311 | show ?thesis | |
| 60567 | 312 | using neq False gz | 
| 60538 | 313 | of_int_div [where ?'a = 'a, OF gz gcd_dvd1 [of "a * b' + b * a'" "b * b'"]] | 
| 314 | of_int_div [where ?'a = 'a, OF gz gcd_dvd2 [of "a * b' + b * a'" "b * b'"]] | |
| 315 | by (simp add: x y Nadd_def INum_def normNum_def Let_def) (simp add: field_simps) | |
| 316 | qed | |
| 317 | qed | |
| 24197 | 318 | qed | 
| 319 | ||
| 60538 | 320 | lemma Nmul[simp]: "INum (x *\<^sub>N y) = INum x * (INum y:: 'a::{field_char_0,field})"
 | 
| 321 | proof - | |
| 322 | let ?z = "0::'a" | |
| 323 | obtain a b where x: "x = (a, b)" by (cases x) | |
| 324 | obtain a' b' where y: "y = (a', b')" by (cases y) | |
| 325 | consider "a = 0 \<or> a' = 0 \<or> b = 0 \<or> b' = 0" | "a \<noteq> 0" "a' \<noteq> 0" "b \<noteq> 0" "b' \<noteq> 0" | |
| 326 | by blast | |
| 327 | then show ?thesis | |
| 328 | proof cases | |
| 329 | case 1 | |
| 330 | then show ?thesis | |
| 60698 | 331 | by (auto simp add: x y Nmul_def INum_def) | 
| 60538 | 332 | next | 
| 60567 | 333 | case neq: 2 | 
| 60538 | 334 | let ?g = "gcd (a * a') (b * b')" | 
| 335 | have gz: "?g \<noteq> 0" | |
| 60567 | 336 | using neq by simp | 
| 337 | from neq of_int_div [where ?'a = 'a, OF gz gcd_dvd1 [of "a * a'" "b * b'"]] | |
| 60538 | 338 | of_int_div [where ?'a = 'a , OF gz gcd_dvd2 [of "a * a'" "b * b'"]] | 
| 339 | show ?thesis | |
| 340 | by (simp add: Nmul_def x y Let_def INum_def) | |
| 341 | qed | |
| 342 | qed | |
| 343 | ||
| 60698 | 344 | 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 | 345 | by (simp add: Nneg_def split_def INum_def) | 
| 24197 | 346 | |
| 60698 | 347 | lemma Nsub[simp]: "INum (x -\<^sub>N y) = INum x - (INum y:: 'a::{field_char_0,field})"
 | 
| 44779 | 348 | by (simp add: Nsub_def split_def) | 
| 24197 | 349 | |
| 60698 | 350 | 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 | 351 | by (simp add: Ninv_def INum_def split_def) | 
| 24197 | 352 | |
| 60698 | 353 | lemma Ndiv[simp]: "INum (x \<div>\<^sub>N y) = INum x / (INum y :: 'a::{field_char_0,field})"
 | 
| 44779 | 354 | by (simp add: Ndiv_def) | 
| 24197 | 355 | |
| 44779 | 356 | lemma Nlt0_iff[simp]: | 
| 44780 | 357 | assumes nx: "isnormNum x" | 
| 60698 | 358 |   shows "((INum x :: 'a::{field_char_0,linordered_field}) < 0) = 0>\<^sub>N x"
 | 
| 44779 | 359 | proof - | 
| 44780 | 360 | obtain a b where x: "x = (a, b)" by (cases x) | 
| 60538 | 361 | show ?thesis | 
| 362 | proof (cases "a = 0") | |
| 363 | case True | |
| 364 | then show ?thesis | |
| 365 | by (simp add: x Nlt0_def INum_def) | |
| 366 | next | |
| 367 | case False | |
| 368 | then have b: "(of_int b::'a) > 0" | |
| 44780 | 369 | using nx by (simp add: x isnormNum_def) | 
| 24197 | 370 | from pos_divide_less_eq[OF b, where b="of_int a" and a="0::'a"] | 
| 60538 | 371 | show ?thesis | 
| 372 | by (simp add: x Nlt0_def INum_def) | |
| 373 | qed | |
| 24197 | 374 | qed | 
| 375 | ||
| 44779 | 376 | lemma Nle0_iff[simp]: | 
| 377 | assumes nx: "isnormNum x" | |
| 60538 | 378 |   shows "((INum x :: 'a::{field_char_0,linordered_field}) \<le> 0) = 0\<ge>\<^sub>N x"
 | 
| 44779 | 379 | proof - | 
| 44780 | 380 | obtain a b where x: "x = (a, b)" by (cases x) | 
| 60538 | 381 | show ?thesis | 
| 382 | proof (cases "a = 0") | |
| 383 | case True | |
| 384 | then show ?thesis | |
| 385 | by (simp add: x Nle0_def INum_def) | |
| 386 | next | |
| 387 | case False | |
| 388 | then have b: "(of_int b :: 'a) > 0" | |
| 44780 | 389 | using nx by (simp add: x isnormNum_def) | 
| 24197 | 390 | from pos_divide_le_eq[OF b, where b="of_int a" and a="0::'a"] | 
| 60538 | 391 | show ?thesis | 
| 392 | by (simp add: x Nle0_def INum_def) | |
| 393 | qed | |
| 24197 | 394 | qed | 
| 395 | ||
| 44779 | 396 | lemma Ngt0_iff[simp]: | 
| 397 | assumes nx: "isnormNum x" | |
| 60698 | 398 |   shows "((INum x :: 'a::{field_char_0,linordered_field}) > 0) = 0<\<^sub>N x"
 | 
| 44779 | 399 | proof - | 
| 44780 | 400 | obtain a b where x: "x = (a, b)" by (cases x) | 
| 60538 | 401 | show ?thesis | 
| 402 | proof (cases "a = 0") | |
| 403 | case True | |
| 404 | then show ?thesis | |
| 405 | by (simp add: x Ngt0_def INum_def) | |
| 406 | next | |
| 407 | case False | |
| 408 | then have b: "(of_int b::'a) > 0" | |
| 409 | using nx by (simp add: x isnormNum_def) | |
| 24197 | 410 | from pos_less_divide_eq[OF b, where b="of_int a" and a="0::'a"] | 
| 60538 | 411 | show ?thesis | 
| 412 | by (simp add: x Ngt0_def INum_def) | |
| 413 | qed | |
| 24197 | 414 | qed | 
| 415 | ||
| 44779 | 416 | lemma Nge0_iff[simp]: | 
| 417 | assumes nx: "isnormNum x" | |
| 60698 | 418 |   shows "(INum x :: 'a::{field_char_0,linordered_field}) \<ge> 0 \<longleftrightarrow> 0\<le>\<^sub>N x"
 | 
| 44779 | 419 | proof - | 
| 44780 | 420 | obtain a b where x: "x = (a, b)" by (cases x) | 
| 60538 | 421 | show ?thesis | 
| 422 | proof (cases "a = 0") | |
| 423 | case True | |
| 424 | then show ?thesis | |
| 425 | by (simp add: x Nge0_def INum_def) | |
| 426 | next | |
| 427 | case False | |
| 428 | then have b: "(of_int b::'a) > 0" | |
| 429 | using nx by (simp add: x isnormNum_def) | |
| 44779 | 430 | from pos_le_divide_eq[OF b, where b="of_int a" and a="0::'a"] | 
| 60538 | 431 | show ?thesis | 
| 432 | by (simp add: x Nge0_def INum_def) | |
| 433 | qed | |
| 44779 | 434 | qed | 
| 435 | ||
| 436 | lemma Nlt_iff[simp]: | |
| 60538 | 437 | assumes nx: "isnormNum x" | 
| 438 | and ny: "isnormNum y" | |
| 60698 | 439 |   shows "((INum x :: 'a::{field_char_0,linordered_field}) < INum y) \<longleftrightarrow> x <\<^sub>N y"
 | 
| 44779 | 440 | proof - | 
| 24197 | 441 | let ?z = "0::'a" | 
| 60698 | 442 | have "((INum x ::'a) < INum y) \<longleftrightarrow> INum (x -\<^sub>N y) < ?z" | 
| 44779 | 443 | using nx ny by simp | 
| 60698 | 444 | also have "\<dots> \<longleftrightarrow> 0>\<^sub>N (x -\<^sub>N y)" | 
| 44779 | 445 | using Nlt0_iff[OF Nsub_normN[OF ny]] by simp | 
| 60538 | 446 | finally show ?thesis | 
| 447 | by (simp add: Nlt_def) | |
| 24197 | 448 | qed | 
| 449 | ||
| 44779 | 450 | lemma Nle_iff[simp]: | 
| 60538 | 451 | assumes nx: "isnormNum x" | 
| 452 | and ny: "isnormNum y" | |
| 60698 | 453 |   shows "((INum x :: 'a::{field_char_0,linordered_field}) \<le> INum y) \<longleftrightarrow> x \<le>\<^sub>N y"
 | 
| 44779 | 454 | proof - | 
| 60698 | 455 | have "((INum x ::'a) \<le> INum y) \<longleftrightarrow> INum (x -\<^sub>N y) \<le> (0::'a)" | 
| 44779 | 456 | using nx ny by simp | 
| 60698 | 457 | also have "\<dots> \<longleftrightarrow> 0\<ge>\<^sub>N (x -\<^sub>N y)" | 
| 44779 | 458 | using Nle0_iff[OF Nsub_normN[OF ny]] by simp | 
| 60538 | 459 | finally show ?thesis | 
| 460 | by (simp add: Nle_def) | |
| 24197 | 461 | qed | 
| 462 | ||
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 463 | lemma Nadd_commute: | 
| 60538 | 464 |   assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
 | 
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 465 | shows "x +\<^sub>N y = y +\<^sub>N x" | 
| 44779 | 466 | proof - | 
| 60538 | 467 | have n: "isnormNum (x +\<^sub>N y)" "isnormNum (y +\<^sub>N x)" | 
| 468 | by simp_all | |
| 469 | have "(INum (x +\<^sub>N y)::'a) = INum (y +\<^sub>N x)" | |
| 470 | by simp | |
| 471 | with isnormNum_unique[OF n] show ?thesis | |
| 472 | by simp | |
| 24197 | 473 | qed | 
| 474 | ||
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 475 | lemma [simp]: | 
| 60538 | 476 |   assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
 | 
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 477 | shows "(0, b) +\<^sub>N y = normNum y" | 
| 44780 | 478 | and "(a, 0) +\<^sub>N y = normNum y" | 
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 479 | and "x +\<^sub>N (0, b) = normNum x" | 
| 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 480 | and "x +\<^sub>N (a, 0) = normNum x" | 
| 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 481 | apply (simp add: Nadd_def split_def) | 
| 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 482 | apply (simp add: Nadd_def split_def) | 
| 60538 | 483 | apply (subst Nadd_commute) | 
| 484 | apply (simp add: Nadd_def split_def) | |
| 485 | apply (subst Nadd_commute) | |
| 486 | apply (simp add: Nadd_def split_def) | |
| 24197 | 487 | done | 
| 488 | ||
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 489 | lemma normNum_nilpotent_aux[simp]: | 
| 60538 | 490 |   assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
 | 
| 44780 | 491 | assumes nx: "isnormNum x" | 
| 24197 | 492 | shows "normNum x = x" | 
| 44779 | 493 | proof - | 
| 24197 | 494 | let ?a = "normNum x" | 
| 495 | have n: "isnormNum ?a" by simp | |
| 44779 | 496 | have th: "INum ?a = (INum x ::'a)" by simp | 
| 497 | with isnormNum_unique[OF n nx] show ?thesis by simp | |
| 24197 | 498 | qed | 
| 499 | ||
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 500 | lemma normNum_nilpotent[simp]: | 
| 60538 | 501 |   assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
 | 
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 502 | shows "normNum (normNum x) = normNum x" | 
| 24197 | 503 | by simp | 
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 504 | |
| 60698 | 505 | lemma normNum0[simp]: "normNum (0, b) = 0\<^sub>N" "normNum (a, 0) = 0\<^sub>N" | 
| 24197 | 506 | by (simp_all add: normNum_def) | 
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 507 | |
| 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 508 | lemma normNum_Nadd: | 
| 60538 | 509 |   assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
 | 
| 510 | shows "normNum (x +\<^sub>N y) = x +\<^sub>N y" | |
| 511 | by simp | |
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 512 | |
| 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 513 | lemma Nadd_normNum1[simp]: | 
| 60538 | 514 |   assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
 | 
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 515 | shows "normNum x +\<^sub>N y = x +\<^sub>N y" | 
| 44779 | 516 | proof - | 
| 60698 | 517 | have n: "isnormNum (normNum x +\<^sub>N y)" "isnormNum (x +\<^sub>N y)" | 
| 518 | by simp_all | |
| 519 | have "INum (normNum x +\<^sub>N y) = INum x + (INum y :: 'a)" | |
| 520 | by simp | |
| 521 | also have "\<dots> = INum (x +\<^sub>N y)" | |
| 522 | by simp | |
| 523 | finally show ?thesis | |
| 524 | using isnormNum_unique[OF n] by simp | |
| 24197 | 525 | qed | 
| 526 | ||
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 527 | lemma Nadd_normNum2[simp]: | 
| 60538 | 528 |   assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
 | 
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 529 | shows "x +\<^sub>N normNum y = x +\<^sub>N y" | 
| 44779 | 530 | proof - | 
| 60698 | 531 | have n: "isnormNum (x +\<^sub>N normNum y)" "isnormNum (x +\<^sub>N y)" | 
| 532 | by simp_all | |
| 533 | have "INum (x +\<^sub>N normNum y) = INum x + (INum y :: 'a)" | |
| 534 | by simp | |
| 535 | also have "\<dots> = INum (x +\<^sub>N y)" | |
| 536 | by simp | |
| 537 | finally show ?thesis | |
| 538 | using isnormNum_unique[OF n] by simp | |
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 539 | qed | 
| 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 540 | |
| 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 541 | lemma Nadd_assoc: | 
| 60538 | 542 |   assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
 | 
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 543 | shows "x +\<^sub>N y +\<^sub>N z = x +\<^sub>N (y +\<^sub>N z)" | 
| 44779 | 544 | proof - | 
| 60698 | 545 | have n: "isnormNum (x +\<^sub>N y +\<^sub>N z)" "isnormNum (x +\<^sub>N (y +\<^sub>N z))" | 
| 546 | by simp_all | |
| 547 | have "INum (x +\<^sub>N y +\<^sub>N z) = (INum (x +\<^sub>N (y +\<^sub>N z)) :: 'a)" | |
| 548 | by simp | |
| 549 | with isnormNum_unique[OF n] show ?thesis | |
| 550 | by simp | |
| 24197 | 551 | qed | 
| 552 | ||
| 553 | lemma Nmul_commute: "isnormNum x \<Longrightarrow> isnormNum y \<Longrightarrow> x *\<^sub>N y = y *\<^sub>N x" | |
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
56544diff
changeset | 554 | by (simp add: Nmul_def split_def Let_def gcd_commute_int mult.commute) | 
| 24197 | 555 | |
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 556 | lemma Nmul_assoc: | 
| 60538 | 557 |   assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
 | 
| 558 | assumes nx: "isnormNum x" | |
| 559 | and ny: "isnormNum y" | |
| 560 | and nz: "isnormNum z" | |
| 24197 | 561 | shows "x *\<^sub>N y *\<^sub>N z = x *\<^sub>N (y *\<^sub>N z)" | 
| 44779 | 562 | proof - | 
| 44780 | 563 | from nx ny nz have n: "isnormNum (x *\<^sub>N y *\<^sub>N z)" "isnormNum (x *\<^sub>N (y *\<^sub>N z))" | 
| 24197 | 564 | by simp_all | 
| 60698 | 565 | have "INum (x +\<^sub>N y +\<^sub>N z) = (INum (x +\<^sub>N (y +\<^sub>N z)) :: 'a)" | 
| 566 | by simp | |
| 567 | with isnormNum_unique[OF n] show ?thesis | |
| 568 | by simp | |
| 24197 | 569 | qed | 
| 570 | ||
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 571 | lemma Nsub0: | 
| 60538 | 572 |   assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
 | 
| 573 | assumes x: "isnormNum x" | |
| 574 | and y: "isnormNum y" | |
| 44780 | 575 | shows "x -\<^sub>N y = 0\<^sub>N \<longleftrightarrow> x = y" | 
| 44779 | 576 | proof - | 
| 44780 | 577 | from isnormNum_unique[where 'a = 'a, OF Nsub_normN[OF y], where y="0\<^sub>N"] | 
| 60698 | 578 | have "x -\<^sub>N y = 0\<^sub>N \<longleftrightarrow> INum (x -\<^sub>N y) = (INum 0\<^sub>N :: 'a)" | 
| 579 | by simp | |
| 580 | also have "\<dots> \<longleftrightarrow> INum x = (INum y :: 'a)" | |
| 581 | by simp | |
| 582 | also have "\<dots> \<longleftrightarrow> x = y" | |
| 583 | using x y by simp | |
| 44779 | 584 | finally show ?thesis . | 
| 24197 | 585 | qed | 
| 586 | ||
| 587 | lemma Nmul0[simp]: "c *\<^sub>N 0\<^sub>N = 0\<^sub>N" " 0\<^sub>N *\<^sub>N c = 0\<^sub>N" | |
| 588 | by (simp_all add: Nmul_def Let_def split_def) | |
| 589 | ||
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 590 | lemma Nmul_eq0[simp]: | 
| 60538 | 591 |   assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
 | 
| 592 | assumes nx: "isnormNum x" | |
| 593 | and ny: "isnormNum y" | |
| 44780 | 594 | shows "x*\<^sub>N y = 0\<^sub>N \<longleftrightarrow> x = 0\<^sub>N \<or> y = 0\<^sub>N" | 
| 44779 | 595 | proof - | 
| 44780 | 596 | obtain a b where x: "x = (a, b)" by (cases x) | 
| 597 | obtain a' b' where y: "y = (a', b')" by (cases y) | |
| 44779 | 598 | have n0: "isnormNum 0\<^sub>N" by simp | 
| 44780 | 599 | show ?thesis using nx ny | 
| 44779 | 600 | apply (simp only: isnormNum_unique[where ?'a = 'a, OF Nmul_normN[OF nx ny] n0, symmetric] | 
| 601 | Nmul[where ?'a = 'a]) | |
| 44780 | 602 | apply (simp add: x y INum_def split_def isnormNum_def split: split_if_asm) | 
| 44779 | 603 | done | 
| 24197 | 604 | qed | 
| 44779 | 605 | |
| 24197 | 606 | lemma Nneg_Nneg[simp]: "~\<^sub>N (~\<^sub>N c) = c" | 
| 607 | by (simp add: Nneg_def split_def) | |
| 608 | ||
| 60538 | 609 | lemma Nmul1[simp]: "isnormNum c \<Longrightarrow> (1)\<^sub>N *\<^sub>N c = c" "isnormNum c \<Longrightarrow> c *\<^sub>N (1)\<^sub>N = c" | 
| 24197 | 610 | 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 | 611 | 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 | 612 | done | 
| 24197 | 613 | |
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 614 | end |