| author | haftmann | 
| Mon, 14 Apr 2025 20:19:05 +0200 | |
| changeset 82509 | c476149a3790 | 
| parent 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)" | |
| 80105 
2fa018321400
Streamlining of many more archaic proofs
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 143 | by (simp add: Ninv_def isnormNum_def split_def gcd.commute split: if_split_asm) | 
| 24197 | 144 | |
| 60538 | 145 | lemma isnormNum_int[simp]: "isnormNum 0\<^sub>N" "isnormNum ((1::int)\<^sub>N)" "i \<noteq> 0 \<Longrightarrow> isnormNum (i)\<^sub>N" | 
| 31706 | 146 | by (simp_all add: isnormNum_def) | 
| 24197 | 147 | |
| 148 | ||
| 60533 | 149 | text \<open>Relations over Num\<close> | 
| 24197 | 150 | |
| 69597 | 151 | definition Nlt0:: "Num \<Rightarrow> bool" (\<open>0>\<^sub>N\<close>) | 
| 60538 | 152 | where "Nlt0 = (\<lambda>(a, b). a < 0)" | 
| 24197 | 153 | |
| 69597 | 154 | definition Nle0:: "Num \<Rightarrow> bool" (\<open>0\<ge>\<^sub>N\<close>) | 
| 60538 | 155 | where "Nle0 = (\<lambda>(a, b). a \<le> 0)" | 
| 24197 | 156 | |
| 69597 | 157 | definition Ngt0:: "Num \<Rightarrow> bool" (\<open>0<\<^sub>N\<close>) | 
| 60538 | 158 | where "Ngt0 = (\<lambda>(a, b). a > 0)" | 
| 24197 | 159 | |
| 69597 | 160 | definition Nge0:: "Num \<Rightarrow> bool" (\<open>0\<le>\<^sub>N\<close>) | 
| 60538 | 161 | where "Nge0 = (\<lambda>(a, b). a \<ge> 0)" | 
| 24197 | 162 | |
| 69597 | 163 | definition Nlt :: "Num \<Rightarrow> Num \<Rightarrow> bool" (infix \<open><\<^sub>N\<close> 55) | 
| 44779 | 164 | where "Nlt = (\<lambda>a b. 0>\<^sub>N (a -\<^sub>N b))" | 
| 24197 | 165 | |
| 69597 | 166 | definition Nle :: "Num \<Rightarrow> Num \<Rightarrow> bool" (infix \<open>\<le>\<^sub>N\<close> 55) | 
| 44779 | 167 | where "Nle = (\<lambda>a b. 0\<ge>\<^sub>N (a -\<^sub>N b))" | 
| 24197 | 168 | |
| 60538 | 169 | definition "INum = (\<lambda>(a, b). of_int a / of_int b)" | 
| 24197 | 170 | |
| 60538 | 171 | lemma INum_int [simp]: "INum (i)\<^sub>N = (of_int i ::'a::field)" "INum 0\<^sub>N = (0::'a::field)" | 
| 24197 | 172 | by (simp_all add: INum_def) | 
| 173 | ||
| 44780 | 174 | lemma isnormNum_unique[simp]: | 
| 60538 | 175 | assumes na: "isnormNum x" | 
| 176 | and nb: "isnormNum y" | |
| 68442 | 177 | shows "(INum x ::'a::field_char_0) = INum y \<longleftrightarrow> x = y" | 
| 60538 | 178 | (is "?lhs = ?rhs") | 
| 24197 | 179 | proof | 
| 44780 | 180 | obtain a b where x: "x = (a, b)" by (cases x) | 
| 181 | obtain a' b' where y: "y = (a', b')" by (cases y) | |
| 60538 | 182 | consider "a = 0 \<or> b = 0 \<or> a' = 0 \<or> b' = 0" | "a \<noteq> 0" "b \<noteq> 0" "a' \<noteq> 0" "b' \<noteq> 0" | 
| 183 | by blast | |
| 184 | then show ?rhs if H: ?lhs | |
| 185 | proof cases | |
| 186 | case 1 | |
| 187 | then show ?thesis | |
| 62390 | 188 | using na nb H by (simp add: x y INum_def split_def isnormNum_def split: if_split_asm) | 
| 60538 | 189 | next | 
| 190 | case 2 | |
| 191 | with na nb have pos: "b > 0" "b' > 0" | |
| 192 | by (simp_all add: x y isnormNum_def) | |
| 60567 | 193 | from H \<open>b \<noteq> 0\<close> \<open>b' \<noteq> 0\<close> have eq: "a * b' = a' * b" | 
| 44780 | 194 | by (simp add: x y INum_def eq_divide_eq divide_eq_eq of_int_mult[symmetric] del: of_int_mult) | 
| 60538 | 195 | from \<open>a \<noteq> 0\<close> \<open>a' \<noteq> 0\<close> na nb | 
| 196 | have gcd1: "gcd a b = 1" "gcd b a = 1" "gcd a' b' = 1" "gcd b' a' = 1" | |
| 62348 | 197 | by (simp_all add: x y isnormNum_def add: gcd.commute) | 
| 67051 | 198 | then have "coprime a b" "coprime b a" "coprime a' b'" "coprime b' a'" | 
| 199 | by (simp_all add: coprime_iff_gcd_eq_1) | |
| 44780 | 200 | from eq have raw_dvd: "a dvd a' * b" "b dvd b' * a" "a' dvd a * b'" "b' dvd b * a'" | 
| 80105 
2fa018321400
Streamlining of many more archaic proofs
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 201 | by (algebra|simp)+ | 
| 67051 | 202 | then have eq1: "b = b'" | 
| 203 | using pos \<open>coprime b a\<close> \<open>coprime b' a'\<close> | |
| 204 | by (auto simp add: coprime_dvd_mult_left_iff intro: associated_eqI) | |
| 205 | with eq have "a = a'" using pos by simp | |
| 206 | with \<open>b = b'\<close> show ?thesis by (simp add: x y) | |
| 60538 | 207 | qed | 
| 208 | show ?lhs if ?rhs | |
| 209 | using that by simp | |
| 24197 | 210 | qed | 
| 211 | ||
| 68442 | 212 | lemma isnormNum0[simp]: "isnormNum x \<Longrightarrow> INum x = (0::'a::field_char_0) \<longleftrightarrow> x = 0\<^sub>N" | 
| 24197 | 213 | unfolding INum_int(2)[symmetric] | 
| 44779 | 214 | by (rule isnormNum_unique) simp_all | 
| 24197 | 215 | |
| 60538 | 216 | lemma of_int_div_aux: | 
| 217 | assumes "d \<noteq> 0" | |
| 218 | shows "(of_int x ::'a::field_char_0) / of_int d = | |
| 219 | of_int (x div d) + (of_int (x mod d)) / of_int d" | |
| 24197 | 220 | proof - | 
| 60538 | 221 | let ?t = "of_int (x div d) * (of_int d ::'a) + of_int (x mod d)" | 
| 24197 | 222 | let ?f = "\<lambda>x. x / of_int d" | 
| 223 | have "x = (x div d) * d + x mod d" | |
| 224 | by auto | |
| 225 | then have eq: "of_int x = ?t" | |
| 226 | by (simp only: of_int_mult[symmetric] of_int_add [symmetric]) | |
| 44780 | 227 | then have "of_int x / of_int d = ?t / of_int d" | 
| 24197 | 228 | using cong[OF refl[of ?f] eq] by simp | 
| 60538 | 229 | then show ?thesis | 
| 230 | by (simp add: add_divide_distrib algebra_simps \<open>d \<noteq> 0\<close>) | |
| 24197 | 231 | qed | 
| 232 | ||
| 60538 | 233 | lemma of_int_div: | 
| 234 | fixes d :: int | |
| 235 | assumes "d \<noteq> 0" "d dvd n" | |
| 236 | shows "(of_int (n div d) ::'a::field_char_0) = of_int n / of_int d" | |
| 237 | using assms of_int_div_aux [of d n, where ?'a = 'a] by simp | |
| 24197 | 238 | |
| 68442 | 239 | lemma normNum[simp]: "INum (normNum x) = (INum x :: 'a::field_char_0)" | 
| 44779 | 240 | proof - | 
| 44780 | 241 | obtain a b where x: "x = (a, b)" by (cases x) | 
| 60538 | 242 | consider "a = 0 \<or> b = 0" | "a \<noteq> 0" "b \<noteq> 0" by blast | 
| 243 | then show ?thesis | |
| 244 | proof cases | |
| 245 | case 1 | |
| 246 | then show ?thesis | |
| 247 | by (simp add: x INum_def normNum_def split_def Let_def) | |
| 248 | next | |
| 60567 | 249 | case ab: 2 | 
| 31706 | 250 | let ?g = "gcd a b" | 
| 60567 | 251 | from ab have g: "?g \<noteq> 0"by simp | 
| 60538 | 252 | from of_int_div[OF g, where ?'a = 'a] show ?thesis | 
| 253 | by (auto simp add: x INum_def normNum_def split_def Let_def) | |
| 254 | qed | |
| 24197 | 255 | qed | 
| 256 | ||
| 68442 | 257 | lemma INum_normNum_iff: "(INum x ::'a::field_char_0) = INum y \<longleftrightarrow> normNum x = normNum y" | 
| 60538 | 258 | (is "?lhs \<longleftrightarrow> _") | 
| 24197 | 259 | proof - | 
| 260 | have "normNum x = normNum y \<longleftrightarrow> (INum (normNum x) :: 'a) = INum (normNum y)" | |
| 261 | by (simp del: normNum) | |
| 262 | also have "\<dots> = ?lhs" by simp | |
| 263 | finally show ?thesis by simp | |
| 264 | qed | |
| 265 | ||
| 68442 | 266 | lemma Nadd[simp]: "INum (x +\<^sub>N y) = INum x + (INum y :: 'a :: field_char_0)" | 
| 44779 | 267 | proof - | 
| 24197 | 268 | let ?z = "0::'a" | 
| 44780 | 269 | obtain a b where x: "x = (a, b)" by (cases x) | 
| 270 | obtain a' b' where y: "y = (a', b')" by (cases y) | |
| 60538 | 271 | consider "a = 0 \<or> a'= 0 \<or> b =0 \<or> b' = 0" | "a \<noteq> 0" "a'\<noteq> 0" "b \<noteq> 0" "b' \<noteq> 0" | 
| 272 | by blast | |
| 273 | then show ?thesis | |
| 274 | proof cases | |
| 275 | case 1 | |
| 276 | then show ?thesis | |
| 80105 
2fa018321400
Streamlining of many more archaic proofs
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 277 | by (auto simp: x y INum_def Nadd_def normNum_def Let_def of_int_div) | 
| 60538 | 278 | next | 
| 60567 | 279 | case neq: 2 | 
| 60538 | 280 | show ?thesis | 
| 281 | proof (cases "a * b' + b * a' = 0") | |
| 282 | case True | |
| 283 | then have "of_int (a * b' + b * a') / (of_int b * of_int b') = ?z" | |
| 284 | by simp | |
| 285 | then have "of_int b' * of_int a / (of_int b * of_int b') + | |
| 286 | of_int b * of_int a' / (of_int b * of_int b') = ?z" | |
| 287 | by (simp add: add_divide_distrib) | |
| 288 | then have th: "of_int a / of_int b + of_int a' / of_int b' = ?z" | |
| 60567 | 289 | using neq by simp | 
| 290 | from True neq show ?thesis | |
| 60538 | 291 | by (simp add: x y th Nadd_def normNum_def INum_def split_def) | 
| 292 | next | |
| 293 | case False | |
| 294 | let ?g = "gcd (a * b' + b * a') (b * b')" | |
| 295 | have gz: "?g \<noteq> 0" | |
| 296 | using False by simp | |
| 297 | show ?thesis | |
| 60567 | 298 | using neq False gz | 
| 60538 | 299 | of_int_div [where ?'a = 'a, OF gz gcd_dvd1 [of "a * b' + b * a'" "b * b'"]] | 
| 300 | of_int_div [where ?'a = 'a, OF gz gcd_dvd2 [of "a * b' + b * a'" "b * b'"]] | |
| 301 | by (simp add: x y Nadd_def INum_def normNum_def Let_def) (simp add: field_simps) | |
| 302 | qed | |
| 303 | qed | |
| 24197 | 304 | qed | 
| 305 | ||
| 68442 | 306 | lemma Nmul[simp]: "INum (x *\<^sub>N y) = INum x * (INum y:: 'a::field_char_0)" | 
| 60538 | 307 | proof - | 
| 308 | let ?z = "0::'a" | |
| 309 | obtain a b where x: "x = (a, b)" by (cases x) | |
| 310 | obtain a' b' where y: "y = (a', b')" by (cases y) | |
| 311 | consider "a = 0 \<or> a' = 0 \<or> b = 0 \<or> b' = 0" | "a \<noteq> 0" "a' \<noteq> 0" "b \<noteq> 0" "b' \<noteq> 0" | |
| 312 | by blast | |
| 313 | then show ?thesis | |
| 314 | proof cases | |
| 315 | case 1 | |
| 316 | then show ?thesis | |
| 60698 | 317 | by (auto simp add: x y Nmul_def INum_def) | 
| 60538 | 318 | next | 
| 60567 | 319 | case neq: 2 | 
| 60538 | 320 | let ?g = "gcd (a * a') (b * b')" | 
| 321 | have gz: "?g \<noteq> 0" | |
| 60567 | 322 | using neq by simp | 
| 323 | from neq of_int_div [where ?'a = 'a, OF gz gcd_dvd1 [of "a * a'" "b * b'"]] | |
| 60538 | 324 | of_int_div [where ?'a = 'a , OF gz gcd_dvd2 [of "a * a'" "b * b'"]] | 
| 325 | show ?thesis | |
| 326 | by (simp add: Nmul_def x y Let_def INum_def) | |
| 327 | qed | |
| 328 | qed | |
| 329 | ||
| 60698 | 330 | 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 | 331 | by (simp add: Nneg_def split_def INum_def) | 
| 24197 | 332 | |
| 68442 | 333 | lemma Nsub[simp]: "INum (x -\<^sub>N y) = INum x - (INum y:: 'a::field_char_0)" | 
| 44779 | 334 | by (simp add: Nsub_def split_def) | 
| 24197 | 335 | |
| 60698 | 336 | 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 | 337 | by (simp add: Ninv_def INum_def split_def) | 
| 24197 | 338 | |
| 68442 | 339 | lemma Ndiv[simp]: "INum (x \<div>\<^sub>N y) = INum x / (INum y :: 'a::field_char_0)" | 
| 44779 | 340 | by (simp add: Ndiv_def) | 
| 24197 | 341 | |
| 44779 | 342 | lemma Nlt0_iff[simp]: | 
| 44780 | 343 | assumes nx: "isnormNum x" | 
| 60698 | 344 |   shows "((INum x :: 'a::{field_char_0,linordered_field}) < 0) = 0>\<^sub>N x"
 | 
| 44779 | 345 | proof - | 
| 44780 | 346 | obtain a b where x: "x = (a, b)" by (cases x) | 
| 60538 | 347 | show ?thesis | 
| 348 | proof (cases "a = 0") | |
| 349 | case True | |
| 350 | then show ?thesis | |
| 351 | by (simp add: x Nlt0_def INum_def) | |
| 352 | next | |
| 353 | case False | |
| 354 | then have b: "(of_int b::'a) > 0" | |
| 44780 | 355 | using nx by (simp add: x isnormNum_def) | 
| 24197 | 356 | from pos_divide_less_eq[OF b, where b="of_int a" and a="0::'a"] | 
| 60538 | 357 | show ?thesis | 
| 358 | by (simp add: x Nlt0_def INum_def) | |
| 359 | qed | |
| 24197 | 360 | qed | 
| 361 | ||
| 44779 | 362 | lemma Nle0_iff[simp]: | 
| 363 | assumes nx: "isnormNum x" | |
| 60538 | 364 |   shows "((INum x :: 'a::{field_char_0,linordered_field}) \<le> 0) = 0\<ge>\<^sub>N x"
 | 
| 44779 | 365 | proof - | 
| 44780 | 366 | obtain a b where x: "x = (a, b)" by (cases x) | 
| 60538 | 367 | show ?thesis | 
| 368 | proof (cases "a = 0") | |
| 369 | case True | |
| 370 | then show ?thesis | |
| 371 | by (simp add: x Nle0_def INum_def) | |
| 372 | next | |
| 373 | case False | |
| 374 | then have b: "(of_int b :: 'a) > 0" | |
| 44780 | 375 | using nx by (simp add: x isnormNum_def) | 
| 24197 | 376 | from pos_divide_le_eq[OF b, where b="of_int a" and a="0::'a"] | 
| 60538 | 377 | show ?thesis | 
| 378 | by (simp add: x Nle0_def INum_def) | |
| 379 | qed | |
| 24197 | 380 | qed | 
| 381 | ||
| 44779 | 382 | lemma Ngt0_iff[simp]: | 
| 383 | assumes nx: "isnormNum x" | |
| 60698 | 384 |   shows "((INum x :: 'a::{field_char_0,linordered_field}) > 0) = 0<\<^sub>N x"
 | 
| 44779 | 385 | proof - | 
| 44780 | 386 | obtain a b where x: "x = (a, b)" by (cases x) | 
| 60538 | 387 | show ?thesis | 
| 388 | proof (cases "a = 0") | |
| 389 | case True | |
| 390 | then show ?thesis | |
| 391 | by (simp add: x Ngt0_def INum_def) | |
| 392 | next | |
| 393 | case False | |
| 394 | then have b: "(of_int b::'a) > 0" | |
| 395 | using nx by (simp add: x isnormNum_def) | |
| 24197 | 396 | from pos_less_divide_eq[OF b, where b="of_int a" and a="0::'a"] | 
| 60538 | 397 | show ?thesis | 
| 398 | by (simp add: x Ngt0_def INum_def) | |
| 399 | qed | |
| 24197 | 400 | qed | 
| 401 | ||
| 44779 | 402 | lemma Nge0_iff[simp]: | 
| 403 | assumes nx: "isnormNum x" | |
| 60698 | 404 |   shows "(INum x :: 'a::{field_char_0,linordered_field}) \<ge> 0 \<longleftrightarrow> 0\<le>\<^sub>N x"
 | 
| 44779 | 405 | proof - | 
| 44780 | 406 | obtain a b where x: "x = (a, b)" by (cases x) | 
| 60538 | 407 | show ?thesis | 
| 408 | proof (cases "a = 0") | |
| 409 | case True | |
| 410 | then show ?thesis | |
| 411 | by (simp add: x Nge0_def INum_def) | |
| 412 | next | |
| 413 | case False | |
| 414 | then have b: "(of_int b::'a) > 0" | |
| 415 | using nx by (simp add: x isnormNum_def) | |
| 44779 | 416 | from pos_le_divide_eq[OF b, where b="of_int a" and a="0::'a"] | 
| 60538 | 417 | show ?thesis | 
| 418 | by (simp add: x Nge0_def INum_def) | |
| 419 | qed | |
| 44779 | 420 | qed | 
| 421 | ||
| 422 | lemma Nlt_iff[simp]: | |
| 60538 | 423 | assumes nx: "isnormNum x" | 
| 424 | and ny: "isnormNum y" | |
| 60698 | 425 |   shows "((INum x :: 'a::{field_char_0,linordered_field}) < INum y) \<longleftrightarrow> x <\<^sub>N y"
 | 
| 44779 | 426 | proof - | 
| 24197 | 427 | let ?z = "0::'a" | 
| 60698 | 428 | have "((INum x ::'a) < INum y) \<longleftrightarrow> INum (x -\<^sub>N y) < ?z" | 
| 44779 | 429 | using nx ny by simp | 
| 60698 | 430 | also have "\<dots> \<longleftrightarrow> 0>\<^sub>N (x -\<^sub>N y)" | 
| 44779 | 431 | using Nlt0_iff[OF Nsub_normN[OF ny]] by simp | 
| 60538 | 432 | finally show ?thesis | 
| 433 | by (simp add: Nlt_def) | |
| 24197 | 434 | qed | 
| 435 | ||
| 44779 | 436 | lemma Nle_iff[simp]: | 
| 60538 | 437 | assumes nx: "isnormNum x" | 
| 438 | and ny: "isnormNum y" | |
| 60698 | 439 |   shows "((INum x :: 'a::{field_char_0,linordered_field}) \<le> INum y) \<longleftrightarrow> x \<le>\<^sub>N y"
 | 
| 44779 | 440 | proof - | 
| 60698 | 441 | have "((INum x ::'a) \<le> INum y) \<longleftrightarrow> INum (x -\<^sub>N y) \<le> (0::'a)" | 
| 44779 | 442 | using nx ny by simp | 
| 60698 | 443 | also have "\<dots> \<longleftrightarrow> 0\<ge>\<^sub>N (x -\<^sub>N y)" | 
| 44779 | 444 | using Nle0_iff[OF Nsub_normN[OF ny]] by simp | 
| 60538 | 445 | finally show ?thesis | 
| 446 | by (simp add: Nle_def) | |
| 24197 | 447 | qed | 
| 448 | ||
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 449 | lemma Nadd_commute: | 
| 68442 | 450 |   assumes "SORT_CONSTRAINT('a::field_char_0)"
 | 
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 451 | shows "x +\<^sub>N y = y +\<^sub>N x" | 
| 44779 | 452 | proof - | 
| 60538 | 453 | have n: "isnormNum (x +\<^sub>N y)" "isnormNum (y +\<^sub>N x)" | 
| 454 | by simp_all | |
| 455 | have "(INum (x +\<^sub>N y)::'a) = INum (y +\<^sub>N x)" | |
| 456 | by simp | |
| 457 | with isnormNum_unique[OF n] show ?thesis | |
| 458 | by simp | |
| 24197 | 459 | qed | 
| 460 | ||
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 461 | lemma [simp]: | 
| 68442 | 462 |   assumes "SORT_CONSTRAINT('a::field_char_0)"
 | 
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 463 | shows "(0, b) +\<^sub>N y = normNum y" | 
| 44780 | 464 | and "(a, 0) +\<^sub>N y = normNum y" | 
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 465 | and "x +\<^sub>N (0, b) = normNum x" | 
| 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 466 | and "x +\<^sub>N (a, 0) = normNum x" | 
| 80105 
2fa018321400
Streamlining of many more archaic proofs
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 467 | by (simp_all add: Nadd_def normNum_def split_def) | 
| 24197 | 468 | |
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 469 | lemma normNum_nilpotent_aux[simp]: | 
| 68442 | 470 |   assumes "SORT_CONSTRAINT('a::field_char_0)"
 | 
| 44780 | 471 | assumes nx: "isnormNum x" | 
| 24197 | 472 | shows "normNum x = x" | 
| 44779 | 473 | proof - | 
| 24197 | 474 | let ?a = "normNum x" | 
| 475 | have n: "isnormNum ?a" by simp | |
| 44779 | 476 | have th: "INum ?a = (INum x ::'a)" by simp | 
| 477 | with isnormNum_unique[OF n nx] show ?thesis by simp | |
| 24197 | 478 | qed | 
| 479 | ||
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 480 | lemma normNum_nilpotent[simp]: | 
| 68442 | 481 |   assumes "SORT_CONSTRAINT('a::field_char_0)"
 | 
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 482 | shows "normNum (normNum x) = normNum x" | 
| 24197 | 483 | by simp | 
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 484 | |
| 60698 | 485 | lemma normNum0[simp]: "normNum (0, b) = 0\<^sub>N" "normNum (a, 0) = 0\<^sub>N" | 
| 24197 | 486 | by (simp_all add: normNum_def) | 
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 487 | |
| 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 488 | lemma normNum_Nadd: | 
| 68442 | 489 |   assumes "SORT_CONSTRAINT('a::field_char_0)"
 | 
| 60538 | 490 | shows "normNum (x +\<^sub>N y) = x +\<^sub>N y" | 
| 491 | by simp | |
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 492 | |
| 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 493 | lemma Nadd_normNum1[simp]: | 
| 68442 | 494 |   assumes "SORT_CONSTRAINT('a::field_char_0)"
 | 
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 495 | shows "normNum x +\<^sub>N y = x +\<^sub>N y" | 
| 44779 | 496 | proof - | 
| 60698 | 497 | have n: "isnormNum (normNum x +\<^sub>N y)" "isnormNum (x +\<^sub>N y)" | 
| 498 | by simp_all | |
| 499 | have "INum (normNum x +\<^sub>N y) = INum x + (INum y :: 'a)" | |
| 500 | by simp | |
| 501 | also have "\<dots> = INum (x +\<^sub>N y)" | |
| 502 | by simp | |
| 503 | finally show ?thesis | |
| 504 | using isnormNum_unique[OF n] by simp | |
| 24197 | 505 | qed | 
| 506 | ||
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 507 | lemma Nadd_normNum2[simp]: | 
| 68442 | 508 |   assumes "SORT_CONSTRAINT('a::field_char_0)"
 | 
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 509 | shows "x +\<^sub>N normNum y = x +\<^sub>N y" | 
| 44779 | 510 | proof - | 
| 60698 | 511 | have n: "isnormNum (x +\<^sub>N normNum y)" "isnormNum (x +\<^sub>N y)" | 
| 512 | by simp_all | |
| 513 | have "INum (x +\<^sub>N normNum y) = INum x + (INum y :: 'a)" | |
| 514 | by simp | |
| 515 | also have "\<dots> = INum (x +\<^sub>N y)" | |
| 516 | by simp | |
| 517 | finally show ?thesis | |
| 518 | using isnormNum_unique[OF n] by simp | |
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 519 | qed | 
| 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 520 | |
| 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 521 | lemma Nadd_assoc: | 
| 68442 | 522 |   assumes "SORT_CONSTRAINT('a::field_char_0)"
 | 
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 523 | shows "x +\<^sub>N y +\<^sub>N z = x +\<^sub>N (y +\<^sub>N z)" | 
| 44779 | 524 | proof - | 
| 60698 | 525 | have n: "isnormNum (x +\<^sub>N y +\<^sub>N z)" "isnormNum (x +\<^sub>N (y +\<^sub>N z))" | 
| 526 | by simp_all | |
| 527 | have "INum (x +\<^sub>N y +\<^sub>N z) = (INum (x +\<^sub>N (y +\<^sub>N z)) :: 'a)" | |
| 528 | by simp | |
| 529 | with isnormNum_unique[OF n] show ?thesis | |
| 530 | by simp | |
| 24197 | 531 | qed | 
| 532 | ||
| 533 | lemma Nmul_commute: "isnormNum x \<Longrightarrow> isnormNum y \<Longrightarrow> x *\<^sub>N y = y *\<^sub>N x" | |
| 62348 | 534 | by (simp add: Nmul_def split_def Let_def gcd.commute mult.commute) | 
| 24197 | 535 | |
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 536 | lemma Nmul_assoc: | 
| 68442 | 537 |   assumes "SORT_CONSTRAINT('a::field_char_0)"
 | 
| 60538 | 538 | assumes nx: "isnormNum x" | 
| 539 | and ny: "isnormNum y" | |
| 540 | and nz: "isnormNum z" | |
| 24197 | 541 | shows "x *\<^sub>N y *\<^sub>N z = x *\<^sub>N (y *\<^sub>N z)" | 
| 44779 | 542 | proof - | 
| 44780 | 543 | from nx ny nz have n: "isnormNum (x *\<^sub>N y *\<^sub>N z)" "isnormNum (x *\<^sub>N (y *\<^sub>N z))" | 
| 24197 | 544 | by simp_all | 
| 60698 | 545 | have "INum (x +\<^sub>N y +\<^sub>N z) = (INum (x +\<^sub>N (y +\<^sub>N z)) :: 'a)" | 
| 546 | by simp | |
| 547 | with isnormNum_unique[OF n] show ?thesis | |
| 548 | by simp | |
| 24197 | 549 | qed | 
| 550 | ||
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 551 | lemma Nsub0: | 
| 68442 | 552 |   assumes "SORT_CONSTRAINT('a::field_char_0)"
 | 
| 60538 | 553 | assumes x: "isnormNum x" | 
| 554 | and y: "isnormNum y" | |
| 44780 | 555 | shows "x -\<^sub>N y = 0\<^sub>N \<longleftrightarrow> x = y" | 
| 44779 | 556 | proof - | 
| 44780 | 557 | from isnormNum_unique[where 'a = 'a, OF Nsub_normN[OF y], where y="0\<^sub>N"] | 
| 60698 | 558 | have "x -\<^sub>N y = 0\<^sub>N \<longleftrightarrow> INum (x -\<^sub>N y) = (INum 0\<^sub>N :: 'a)" | 
| 559 | by simp | |
| 560 | also have "\<dots> \<longleftrightarrow> INum x = (INum y :: 'a)" | |
| 561 | by simp | |
| 562 | also have "\<dots> \<longleftrightarrow> x = y" | |
| 563 | using x y by simp | |
| 44779 | 564 | finally show ?thesis . | 
| 24197 | 565 | qed | 
| 566 | ||
| 567 | lemma Nmul0[simp]: "c *\<^sub>N 0\<^sub>N = 0\<^sub>N" " 0\<^sub>N *\<^sub>N c = 0\<^sub>N" | |
| 568 | by (simp_all add: Nmul_def Let_def split_def) | |
| 569 | ||
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 570 | lemma Nmul_eq0[simp]: | 
| 68442 | 571 |   assumes "SORT_CONSTRAINT('a::field_char_0)"
 | 
| 60538 | 572 | assumes nx: "isnormNum x" | 
| 573 | and ny: "isnormNum y" | |
| 44780 | 574 | shows "x*\<^sub>N y = 0\<^sub>N \<longleftrightarrow> x = 0\<^sub>N \<or> y = 0\<^sub>N" | 
| 44779 | 575 | proof - | 
| 44780 | 576 | obtain a b where x: "x = (a, b)" by (cases x) | 
| 577 | obtain a' b' where y: "y = (a', b')" by (cases y) | |
| 44779 | 578 | have n0: "isnormNum 0\<^sub>N" by simp | 
| 44780 | 579 | show ?thesis using nx ny | 
| 80105 
2fa018321400
Streamlining of many more archaic proofs
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 580 | by (metis (no_types, opaque_lifting) INum_int(2) Nmul Nmul0(1) Nmul0(2) isnormNum0 mult_eq_0_iff) | 
| 24197 | 581 | qed | 
| 44779 | 582 | |
| 24197 | 583 | lemma Nneg_Nneg[simp]: "~\<^sub>N (~\<^sub>N c) = c" | 
| 584 | by (simp add: Nneg_def split_def) | |
| 585 | ||
| 60538 | 586 | lemma Nmul1[simp]: "isnormNum c \<Longrightarrow> (1)\<^sub>N *\<^sub>N c = c" "isnormNum c \<Longrightarrow> c *\<^sub>N (1)\<^sub>N = c" | 
| 80105 
2fa018321400
Streamlining of many more archaic proofs
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 587 | by (simp add: Nmul_def Let_def split_def isnormNum_def, metis div_0 div_by_1 surjective_pairing)+ | 
| 24197 | 588 | |
| 28615 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 wenzelm parents: 
27668diff
changeset | 589 | end |