| author | nipkow | 
| Tue, 27 Aug 2019 17:08:51 +0200 | |
| changeset 70614 | 6a2c982363e9 | 
| parent 69064 | 5840724b1d71 | 
| child 80914 | d97fdabd9e2b | 
| permissions | -rw-r--r-- | 
| 64962 | 1 | (* Title: HOL/Decision_Procs/Algebra_Aux.thy | 
| 2 | Author: Stefan Berghofer | |
| 3 | *) | |
| 4 | ||
| 5 | section \<open>Things that can be added to the Algebra library\<close> | |
| 6 | ||
| 7 | theory Algebra_Aux | |
| 67123 | 8 | imports "HOL-Algebra.Ring" | 
| 64962 | 9 | begin | 
| 10 | ||
| 67123 | 11 | definition of_natural :: "('a, 'm) ring_scheme \<Rightarrow> nat \<Rightarrow> 'a" ("\<guillemotleft>_\<guillemotright>\<^sub>\<nat>\<index>")
 | 
| 67399 | 12 | where "\<guillemotleft>n\<guillemotright>\<^sub>\<nat>\<^bsub>R\<^esub> = ((\<oplus>\<^bsub>R\<^esub>) \<one>\<^bsub>R\<^esub> ^^ n) \<zero>\<^bsub>R\<^esub>" | 
| 64962 | 13 | |
| 67123 | 14 | definition of_integer :: "('a, 'm) ring_scheme \<Rightarrow> int \<Rightarrow> 'a" ("\<guillemotleft>_\<guillemotright>\<index>")
 | 
| 15 | where "\<guillemotleft>i\<guillemotright>\<^bsub>R\<^esub> = (if 0 \<le> i then \<guillemotleft>nat i\<guillemotright>\<^sub>\<nat>\<^bsub>R\<^esub> else \<ominus>\<^bsub>R\<^esub> \<guillemotleft>nat (- i)\<guillemotright>\<^sub>\<nat>\<^bsub>R\<^esub>)" | |
| 64962 | 16 | |
| 67123 | 17 | context ring | 
| 18 | begin | |
| 64962 | 19 | |
| 20 | lemma of_nat_0 [simp]: "\<guillemotleft>0\<guillemotright>\<^sub>\<nat> = \<zero>" | |
| 21 | by (simp add: of_natural_def) | |
| 22 | ||
| 23 | lemma of_nat_Suc [simp]: "\<guillemotleft>Suc n\<guillemotright>\<^sub>\<nat> = \<one> \<oplus> \<guillemotleft>n\<guillemotright>\<^sub>\<nat>" | |
| 24 | by (simp add: of_natural_def) | |
| 25 | ||
| 26 | lemma of_int_0 [simp]: "\<guillemotleft>0\<guillemotright> = \<zero>" | |
| 27 | by (simp add: of_integer_def) | |
| 28 | ||
| 29 | lemma of_nat_closed [simp]: "\<guillemotleft>n\<guillemotright>\<^sub>\<nat> \<in> carrier R" | |
| 30 | by (induct n) simp_all | |
| 31 | ||
| 32 | lemma of_int_closed [simp]: "\<guillemotleft>i\<guillemotright> \<in> carrier R" | |
| 33 | by (simp add: of_integer_def) | |
| 34 | ||
| 35 | lemma of_int_minus [simp]: "\<guillemotleft>- i\<guillemotright> = \<ominus> \<guillemotleft>i\<guillemotright>" | |
| 36 | by (simp add: of_integer_def) | |
| 37 | ||
| 38 | lemma of_nat_add [simp]: "\<guillemotleft>m + n\<guillemotright>\<^sub>\<nat> = \<guillemotleft>m\<guillemotright>\<^sub>\<nat> \<oplus> \<guillemotleft>n\<guillemotright>\<^sub>\<nat>" | |
| 39 | by (induct m) (simp_all add: a_ac) | |
| 40 | ||
| 41 | lemma of_nat_diff [simp]: "n \<le> m \<Longrightarrow> \<guillemotleft>m - n\<guillemotright>\<^sub>\<nat> = \<guillemotleft>m\<guillemotright>\<^sub>\<nat> \<ominus> \<guillemotleft>n\<guillemotright>\<^sub>\<nat>" | |
| 42 | proof (induct m arbitrary: n) | |
| 67123 | 43 | case 0 | 
| 44 | then show ?case by (simp add: minus_eq) | |
| 45 | next | |
| 46 | case Suc': (Suc m) | |
| 64962 | 47 | show ?case | 
| 48 | proof (cases n) | |
| 67123 | 49 | case 0 | 
| 50 | then show ?thesis | |
| 51 | by (simp add: minus_eq) | |
| 52 | next | |
| 64962 | 53 | case (Suc k) | 
| 54 | with Suc' have "\<guillemotleft>Suc m - Suc k\<guillemotright>\<^sub>\<nat> = \<guillemotleft>m\<guillemotright>\<^sub>\<nat> \<ominus> \<guillemotleft>k\<guillemotright>\<^sub>\<nat>" by simp | |
| 55 | also have "\<dots> = \<one> \<oplus> \<ominus> \<one> \<oplus> (\<guillemotleft>m\<guillemotright>\<^sub>\<nat> \<ominus> \<guillemotleft>k\<guillemotright>\<^sub>\<nat>)" | |
| 56 | by (simp add: r_neg) | |
| 57 | also have "\<dots> = \<guillemotleft>Suc m\<guillemotright>\<^sub>\<nat> \<ominus> \<guillemotleft>Suc k\<guillemotright>\<^sub>\<nat>" | |
| 58 | by (simp add: minus_eq minus_add a_ac) | |
| 59 | finally show ?thesis using Suc by simp | |
| 67123 | 60 | qed | 
| 61 | qed | |
| 64962 | 62 | |
| 63 | lemma of_int_add [simp]: "\<guillemotleft>i + j\<guillemotright> = \<guillemotleft>i\<guillemotright> \<oplus> \<guillemotleft>j\<guillemotright>" | |
| 64 | proof (cases "0 \<le> i") | |
| 65 | case True | |
| 66 | show ?thesis | |
| 67 | proof (cases "0 \<le> j") | |
| 68 | case True | |
| 67123 | 69 | with \<open>0 \<le> i\<close> show ?thesis | 
| 70 | by (simp add: of_integer_def nat_add_distrib) | |
| 64962 | 71 | next | 
| 72 | case False | |
| 73 | show ?thesis | |
| 74 | proof (cases "0 \<le> i + j") | |
| 75 | case True | |
| 76 | then have "\<guillemotleft>i + j\<guillemotright> = \<guillemotleft>nat (i - (- j))\<guillemotright>\<^sub>\<nat>" | |
| 77 | by (simp add: of_integer_def) | |
| 64998 | 78 | also from True \<open>0 \<le> i\<close> \<open>\<not> 0 \<le> j\<close> | 
| 64962 | 79 | have "nat (i - (- j)) = nat i - nat (- j)" | 
| 80 | by (simp add: nat_diff_distrib) | |
| 64998 | 81 | finally show ?thesis using True \<open>0 \<le> i\<close> \<open>\<not> 0 \<le> j\<close> | 
| 64962 | 82 | by (simp add: minus_eq of_integer_def) | 
| 83 | next | |
| 84 | case False | |
| 85 | then have "\<guillemotleft>i + j\<guillemotright> = \<ominus> \<guillemotleft>nat (- j - i)\<guillemotright>\<^sub>\<nat>" | |
| 86 | by (simp add: of_integer_def) (simp only: diff_conv_add_uminus add_ac) | |
| 64998 | 87 | also from False \<open>0 \<le> i\<close> \<open>\<not> 0 \<le> j\<close> | 
| 64962 | 88 | have "nat (- j - i) = nat (- j) - nat i" | 
| 89 | by (simp add: nat_diff_distrib) | |
| 64998 | 90 | finally show ?thesis using False \<open>0 \<le> i\<close> \<open>\<not> 0 \<le> j\<close> | 
| 64962 | 91 | by (simp add: minus_eq minus_add a_ac of_integer_def) | 
| 92 | qed | |
| 93 | qed | |
| 94 | next | |
| 95 | case False | |
| 96 | show ?thesis | |
| 97 | proof (cases "0 \<le> j") | |
| 98 | case True | |
| 99 | show ?thesis | |
| 100 | proof (cases "0 \<le> i + j") | |
| 101 | case True | |
| 102 | then have "\<guillemotleft>i + j\<guillemotright> = \<guillemotleft>nat (j - (- i))\<guillemotright>\<^sub>\<nat>" | |
| 103 | by (simp add: of_integer_def add_ac) | |
| 64998 | 104 | also from True \<open>\<not> 0 \<le> i\<close> \<open>0 \<le> j\<close> | 
| 64962 | 105 | have "nat (j - (- i)) = nat j - nat (- i)" | 
| 106 | by (simp add: nat_diff_distrib) | |
| 64998 | 107 | finally show ?thesis using True \<open>\<not> 0 \<le> i\<close> \<open>0 \<le> j\<close> | 
| 64962 | 108 | by (simp add: minus_eq minus_add a_ac of_integer_def) | 
| 109 | next | |
| 110 | case False | |
| 111 | then have "\<guillemotleft>i + j\<guillemotright> = \<ominus> \<guillemotleft>nat (- i - j)\<guillemotright>\<^sub>\<nat>" | |
| 112 | by (simp add: of_integer_def) | |
| 64998 | 113 | also from False \<open>\<not> 0 \<le> i\<close> \<open>0 \<le> j\<close> | 
| 64962 | 114 | have "nat (- i - j) = nat (- i) - nat j" | 
| 115 | by (simp add: nat_diff_distrib) | |
| 64998 | 116 | finally show ?thesis using False \<open>\<not> 0 \<le> i\<close> \<open>0 \<le> j\<close> | 
| 64962 | 117 | by (simp add: minus_eq minus_add of_integer_def) | 
| 118 | qed | |
| 119 | next | |
| 120 | case False | |
| 64998 | 121 | with \<open>\<not> 0 \<le> i\<close> show ?thesis | 
| 64962 | 122 | by (simp add: of_integer_def nat_add_distrib minus_add diff_conv_add_uminus | 
| 67123 | 123 | del: add_uminus_conv_diff uminus_add_conv_diff) | 
| 64962 | 124 | qed | 
| 125 | qed | |
| 126 | ||
| 127 | lemma of_int_diff [simp]: "\<guillemotleft>i - j\<guillemotright> = \<guillemotleft>i\<guillemotright> \<ominus> \<guillemotleft>j\<guillemotright>" | |
| 128 | by (simp only: diff_conv_add_uminus of_int_add) (simp add: minus_eq) | |
| 129 | ||
| 130 | lemma of_nat_mult [simp]: "\<guillemotleft>i * j\<guillemotright>\<^sub>\<nat> = \<guillemotleft>i\<guillemotright>\<^sub>\<nat> \<otimes> \<guillemotleft>j\<guillemotright>\<^sub>\<nat>" | |
| 131 | by (induct i) (simp_all add: l_distr) | |
| 132 | ||
| 133 | lemma of_int_mult [simp]: "\<guillemotleft>i * j\<guillemotright> = \<guillemotleft>i\<guillemotright> \<otimes> \<guillemotleft>j\<guillemotright>" | |
| 134 | proof (cases "0 \<le> i") | |
| 135 | case True | |
| 136 | show ?thesis | |
| 137 | proof (cases "0 \<le> j") | |
| 138 | case True | |
| 64998 | 139 | with \<open>0 \<le> i\<close> show ?thesis | 
| 64962 | 140 | by (simp add: of_integer_def nat_mult_distrib zero_le_mult_iff) | 
| 141 | next | |
| 142 | case False | |
| 64998 | 143 | with \<open>0 \<le> i\<close> show ?thesis | 
| 64962 | 144 | by (simp add: of_integer_def zero_le_mult_iff | 
| 145 | minus_mult_right nat_mult_distrib r_minus | |
| 146 | del: minus_mult_right [symmetric]) | |
| 147 | qed | |
| 148 | next | |
| 149 | case False | |
| 150 | show ?thesis | |
| 151 | proof (cases "0 \<le> j") | |
| 152 | case True | |
| 64998 | 153 | with \<open>\<not> 0 \<le> i\<close> show ?thesis | 
| 64962 | 154 | by (simp add: of_integer_def zero_le_mult_iff | 
| 155 | minus_mult_left nat_mult_distrib l_minus | |
| 156 | del: minus_mult_left [symmetric]) | |
| 157 | next | |
| 158 | case False | |
| 64998 | 159 | with \<open>\<not> 0 \<le> i\<close> show ?thesis | 
| 64962 | 160 | by (simp add: of_integer_def zero_le_mult_iff | 
| 161 | minus_mult_minus [of i j, symmetric] nat_mult_distrib | |
| 162 | l_minus r_minus | |
| 163 | del: minus_mult_minus | |
| 164 | minus_mult_left [symmetric] minus_mult_right [symmetric]) | |
| 165 | qed | |
| 166 | qed | |
| 167 | ||
| 168 | lemma of_int_1 [simp]: "\<guillemotleft>1\<guillemotright> = \<one>" | |
| 169 | by (simp add: of_integer_def) | |
| 170 | ||
| 171 | lemma of_int_2: "\<guillemotleft>2\<guillemotright> = \<one> \<oplus> \<one>" | |
| 172 | by (simp add: of_integer_def numeral_2_eq_2) | |
| 173 | ||
| 174 | lemma minus_0_r [simp]: "x \<in> carrier R \<Longrightarrow> x \<ominus> \<zero> = x" | |
| 175 | by (simp add: minus_eq) | |
| 176 | ||
| 177 | lemma minus_0_l [simp]: "x \<in> carrier R \<Longrightarrow> \<zero> \<ominus> x = \<ominus> x" | |
| 178 | by (simp add: minus_eq) | |
| 179 | ||
| 180 | lemma eq_diff0: | |
| 181 | assumes "x \<in> carrier R" "y \<in> carrier R" | |
| 67123 | 182 | shows "x \<ominus> y = \<zero> \<longleftrightarrow> x = y" | 
| 64962 | 183 | proof | 
| 184 | assume "x \<ominus> y = \<zero>" | |
| 185 | with assms have "x \<oplus> (\<ominus> y \<oplus> y) = y" | |
| 186 | by (simp add: minus_eq a_assoc [symmetric]) | |
| 187 | with assms show "x = y" by (simp add: l_neg) | |
| 188 | next | |
| 189 | assume "x = y" | |
| 190 | with assms show "x \<ominus> y = \<zero>" by (simp add: minus_eq r_neg) | |
| 191 | qed | |
| 192 | ||
| 67341 
df79ef3b3a41
Renamed (^) to [^] in preparation of the move from "op X" to (X)
 nipkow parents: 
67123diff
changeset | 193 | lemma power2_eq_square: "x \<in> carrier R \<Longrightarrow> x [^] (2::nat) = x \<otimes> x" | 
| 64962 | 194 | by (simp add: numeral_eq_Suc) | 
| 195 | ||
| 196 | lemma eq_neg_iff_add_eq_0: | |
| 197 | assumes "x \<in> carrier R" "y \<in> carrier R" | |
| 67123 | 198 | shows "x = \<ominus> y \<longleftrightarrow> x \<oplus> y = \<zero>" | 
| 64962 | 199 | proof | 
| 200 | assume "x = \<ominus> y" | |
| 201 | with assms show "x \<oplus> y = \<zero>" by (simp add: l_neg) | |
| 202 | next | |
| 203 | assume "x \<oplus> y = \<zero>" | |
| 204 | with assms have "x \<oplus> (y \<oplus> \<ominus> y) = \<zero> \<oplus> \<ominus> y" | |
| 205 | by (simp add: a_assoc [symmetric]) | |
| 206 | with assms show "x = \<ominus> y" | |
| 207 | by (simp add: r_neg) | |
| 208 | qed | |
| 209 | ||
| 210 | lemma neg_equal_iff_equal: | |
| 211 | assumes x: "x \<in> carrier R" and y: "y \<in> carrier R" | |
| 67123 | 212 | shows "\<ominus> x = \<ominus> y \<longleftrightarrow> x = y" | 
| 64962 | 213 | proof | 
| 214 | assume "\<ominus> x = \<ominus> y" | |
| 215 | then have "\<ominus> (\<ominus> x) = \<ominus> (\<ominus> y)" by simp | |
| 216 | with x y show "x = y" by simp | |
| 217 | next | |
| 218 | assume "x = y" | |
| 219 | then show "\<ominus> x = \<ominus> y" by simp | |
| 220 | qed | |
| 221 | ||
| 222 | lemma neg_equal_swap: | |
| 223 | assumes x: "x \<in> carrier R" and y: "y \<in> carrier R" | |
| 224 | shows "(\<ominus> x = y) = (x = \<ominus> y)" | |
| 225 | using assms neg_equal_iff_equal [of x "\<ominus> y"] | |
| 226 | by simp | |
| 227 | ||
| 228 | lemma mult2: "x \<in> carrier R \<Longrightarrow> x \<oplus> x = \<guillemotleft>2\<guillemotright> \<otimes> x" | |
| 229 | by (simp add: of_int_2 l_distr) | |
| 230 | ||
| 231 | end | |
| 232 | ||
| 67341 
df79ef3b3a41
Renamed (^) to [^] in preparation of the move from "op X" to (X)
 nipkow parents: 
67123diff
changeset | 233 | lemma (in cring) of_int_power [simp]: "\<guillemotleft>i ^ n\<guillemotright> = \<guillemotleft>i\<guillemotright> [^] n" | 
| 64962 | 234 | by (induct n) (simp_all add: m_ac) | 
| 235 | ||
| 67123 | 236 | definition cring_class_ops :: "'a::comm_ring_1 ring" | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
67399diff
changeset | 237 | where "cring_class_ops \<equiv> \<lparr>carrier = UNIV, mult = (*), one = 1, zero = 0, add = (+)\<rparr>" | 
| 64962 | 238 | |
| 239 | lemma cring_class: "cring cring_class_ops" | |
| 240 | apply unfold_locales | |
| 241 | apply (auto simp add: cring_class_ops_def ring_distribs Units_def) | |
| 242 | apply (rule_tac x="- x" in exI) | |
| 243 | apply simp | |
| 244 | done | |
| 245 | ||
| 246 | lemma carrier_class: "x \<in> carrier cring_class_ops" | |
| 247 | by (simp add: cring_class_ops_def) | |
| 248 | ||
| 249 | lemma zero_class: "\<zero>\<^bsub>cring_class_ops\<^esub> = 0" | |
| 250 | by (simp add: cring_class_ops_def) | |
| 251 | ||
| 252 | lemma one_class: "\<one>\<^bsub>cring_class_ops\<^esub> = 1" | |
| 253 | by (simp add: cring_class_ops_def) | |
| 254 | ||
| 255 | lemma plus_class: "x \<oplus>\<^bsub>cring_class_ops\<^esub> y = x + y" | |
| 256 | by (simp add: cring_class_ops_def) | |
| 257 | ||
| 258 | lemma times_class: "x \<otimes>\<^bsub>cring_class_ops\<^esub> y = x * y" | |
| 259 | by (simp add: cring_class_ops_def) | |
| 260 | ||
| 261 | lemma uminus_class: "\<ominus>\<^bsub>cring_class_ops\<^esub> x = - x" | |
| 262 | apply (simp add: a_inv_def m_inv_def cring_class_ops_def) | |
| 263 | apply (rule the_equality) | |
| 264 | apply (simp_all add: eq_neg_iff_add_eq_0) | |
| 265 | done | |
| 266 | ||
| 267 | lemma minus_class: "x \<ominus>\<^bsub>cring_class_ops\<^esub> y = x - y" | |
| 268 | by (simp add: a_minus_def carrier_class plus_class uminus_class) | |
| 269 | ||
| 67341 
df79ef3b3a41
Renamed (^) to [^] in preparation of the move from "op X" to (X)
 nipkow parents: 
67123diff
changeset | 270 | lemma power_class: "x [^]\<^bsub>cring_class_ops\<^esub> n = x ^ n" | 
| 64962 | 271 | by (induct n) (simp_all add: one_class times_class | 
| 272 | monoid.nat_pow_0 [OF comm_monoid.axioms(1) [OF cring.axioms(2) [OF cring_class]]] | |
| 273 | monoid.nat_pow_Suc [OF comm_monoid.axioms(1) [OF cring.axioms(2) [OF cring_class]]]) | |
| 274 | ||
| 275 | lemma of_nat_class: "\<guillemotleft>n\<guillemotright>\<^sub>\<nat>\<^bsub>cring_class_ops\<^esub> = of_nat n" | |
| 276 | by (induct n) (simp_all add: cring_class_ops_def of_natural_def) | |
| 277 | ||
| 278 | lemma of_int_class: "\<guillemotleft>i\<guillemotright>\<^bsub>cring_class_ops\<^esub> = of_int i" | |
| 279 | by (simp add: of_integer_def of_nat_class uminus_class) | |
| 280 | ||
| 281 | lemmas class_simps = zero_class one_class plus_class minus_class uminus_class | |
| 282 | times_class power_class of_nat_class of_int_class carrier_class | |
| 283 | ||
| 284 | interpretation cring_class: cring "cring_class_ops::'a::comm_ring_1 ring" | |
| 67123 | 285 | rewrites "(\<zero>\<^bsub>cring_class_ops\<^esub>::'a) = 0" | 
| 286 | and "(\<one>\<^bsub>cring_class_ops\<^esub>::'a) = 1" | |
| 287 | and "(x::'a) \<oplus>\<^bsub>cring_class_ops\<^esub> y = x + y" | |
| 288 | and "(x::'a) \<otimes>\<^bsub>cring_class_ops\<^esub> y = x * y" | |
| 289 | and "\<ominus>\<^bsub>cring_class_ops\<^esub> (x::'a) = - x" | |
| 290 | and "(x::'a) \<ominus>\<^bsub>cring_class_ops\<^esub> y = x - y" | |
| 67341 
df79ef3b3a41
Renamed (^) to [^] in preparation of the move from "op X" to (X)
 nipkow parents: 
67123diff
changeset | 291 | and "(x::'a) [^]\<^bsub>cring_class_ops\<^esub> n = x ^ n" | 
| 67123 | 292 | and "\<guillemotleft>n\<guillemotright>\<^sub>\<nat>\<^bsub>cring_class_ops\<^esub> = of_nat n" | 
| 293 | and "((\<guillemotleft>i\<guillemotright>\<^bsub>cring_class_ops\<^esub>)::'a) = of_int i" | |
| 294 | and "(Int.of_int (numeral m)::'a) = numeral m" | |
| 64962 | 295 | by (simp_all add: cring_class class_simps) | 
| 296 | ||
| 297 | lemma (in domain) nat_pow_eq_0_iff [simp]: | |
| 67341 
df79ef3b3a41
Renamed (^) to [^] in preparation of the move from "op X" to (X)
 nipkow parents: 
67123diff
changeset | 298 | "a \<in> carrier R \<Longrightarrow> (a [^] (n::nat) = \<zero>) = (a = \<zero> \<and> n \<noteq> 0)" | 
| 64962 | 299 | by (induct n) (auto simp add: integral_iff) | 
| 300 | ||
| 301 | lemma (in domain) square_eq_iff: | |
| 302 | assumes "x \<in> carrier R" "y \<in> carrier R" | |
| 303 | shows "(x \<otimes> x = y \<otimes> y) = (x = y \<or> x = \<ominus> y)" | |
| 304 | proof | |
| 305 | assume "x \<otimes> x = y \<otimes> y" | |
| 306 | with assms have "(x \<ominus> y) \<otimes> (x \<oplus> y) = x \<otimes> y \<oplus> \<ominus> (x \<otimes> y) \<oplus> (y \<otimes> y \<oplus> \<ominus> (y \<otimes> y))" | |
| 307 | by (simp add: r_distr l_distr minus_eq r_minus m_comm a_ac) | |
| 308 | with assms show "x = y \<or> x = \<ominus> y" | |
| 309 | by (simp add: integral_iff eq_neg_iff_add_eq_0 eq_diff0 r_neg) | |
| 310 | next | |
| 311 | assume "x = y \<or> x = \<ominus> y" | |
| 67123 | 312 | with assms show "x \<otimes> x = y \<otimes> y" | 
| 313 | by (auto simp add: l_minus r_minus) | |
| 64962 | 314 | qed | 
| 315 | ||
| 67123 | 316 | definition m_div :: "('a, 'b) ring_scheme \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<oslash>\<index>" 70)
 | 
| 317 | where "x \<oslash>\<^bsub>G\<^esub> y = (if y = \<zero>\<^bsub>G\<^esub> then \<zero>\<^bsub>G\<^esub> else x \<otimes>\<^bsub>G\<^esub> inv\<^bsub>G\<^esub> y)" | |
| 64962 | 318 | |
| 319 | context field | |
| 320 | begin | |
| 321 | ||
| 322 | lemma inv_closed [simp]: "x \<in> carrier R \<Longrightarrow> x \<noteq> \<zero> \<Longrightarrow> inv x \<in> carrier R" | |
| 323 | by (simp add: field_Units) | |
| 324 | ||
| 325 | lemma l_inv [simp]: "x \<in> carrier R \<Longrightarrow> x \<noteq> \<zero> \<Longrightarrow> inv x \<otimes> x = \<one>" | |
| 326 | by (simp add: field_Units) | |
| 327 | ||
| 328 | lemma r_inv [simp]: "x \<in> carrier R \<Longrightarrow> x \<noteq> \<zero> \<Longrightarrow> x \<otimes> inv x = \<one>" | |
| 329 | by (simp add: field_Units) | |
| 330 | ||
| 331 | lemma inverse_unique: | |
| 332 | assumes a: "a \<in> carrier R" | |
| 67123 | 333 | and b: "b \<in> carrier R" | 
| 334 | and ab: "a \<otimes> b = \<one>" | |
| 64962 | 335 | shows "inv a = b" | 
| 336 | proof - | |
| 67123 | 337 | from ab b have *: "a \<noteq> \<zero>" | 
| 338 | by (cases "a = \<zero>") simp_all | |
| 339 | with a have "inv a \<otimes> (a \<otimes> b) = inv a" | |
| 340 | by (simp add: ab) | |
| 341 | with a b * show ?thesis | |
| 342 | by (simp add: m_assoc [symmetric]) | |
| 64962 | 343 | qed | 
| 344 | ||
| 67123 | 345 | lemma nonzero_inverse_inverse_eq: "a \<in> carrier R \<Longrightarrow> a \<noteq> \<zero> \<Longrightarrow> inv (inv a) = a" | 
| 64962 | 346 | by (rule inverse_unique) simp_all | 
| 347 | ||
| 348 | lemma inv_1 [simp]: "inv \<one> = \<one>" | |
| 349 | by (rule inverse_unique) simp_all | |
| 350 | ||
| 351 | lemma nonzero_inverse_mult_distrib: | |
| 67123 | 352 | assumes "a \<in> carrier R" "b \<in> carrier R" | 
| 353 | and "a \<noteq> \<zero>" "b \<noteq> \<zero>" | |
| 64962 | 354 | shows "inv (a \<otimes> b) = inv b \<otimes> inv a" | 
| 355 | proof - | |
| 67123 | 356 | from assms have "a \<otimes> (b \<otimes> inv b) \<otimes> inv a = \<one>" | 
| 357 | by simp | |
| 358 | with assms have eq: "a \<otimes> b \<otimes> (inv b \<otimes> inv a) = \<one>" | |
| 64962 | 359 | by (simp only: m_assoc m_closed inv_closed assms) | 
| 67123 | 360 | from assms show ?thesis | 
| 361 | using inverse_unique [OF _ _ eq] by simp | |
| 64962 | 362 | qed | 
| 363 | ||
| 364 | lemma nonzero_imp_inverse_nonzero: | |
| 365 | assumes "a \<in> carrier R" and "a \<noteq> \<zero>" | |
| 366 | shows "inv a \<noteq> \<zero>" | |
| 367 | proof | |
| 67123 | 368 | assume *: "inv a = \<zero>" | 
| 369 | from assms have **: "\<one> = a \<otimes> inv a" | |
| 370 | by simp | |
| 371 | also from assms have "\<dots> = \<zero>" by (simp add: *) | |
| 64962 | 372 | finally have "\<one> = \<zero>" . | 
| 67123 | 373 | then show False by (simp add: eq_commute) | 
| 64962 | 374 | qed | 
| 375 | ||
| 376 | lemma nonzero_divide_divide_eq_left: | |
| 377 | "a \<in> carrier R \<Longrightarrow> b \<in> carrier R \<Longrightarrow> c \<in> carrier R \<Longrightarrow> b \<noteq> \<zero> \<Longrightarrow> c \<noteq> \<zero> \<Longrightarrow> | |
| 67123 | 378 | a \<oslash> b \<oslash> c = a \<oslash> (b \<otimes> c)" | 
| 64962 | 379 | by (simp add: m_div_def nonzero_inverse_mult_distrib m_ac integral_iff) | 
| 380 | ||
| 381 | lemma nonzero_times_divide_eq: | |
| 382 | "a \<in> carrier R \<Longrightarrow> b \<in> carrier R \<Longrightarrow> c \<in> carrier R \<Longrightarrow> d \<in> carrier R \<Longrightarrow> | |
| 67123 | 383 | b \<noteq> \<zero> \<Longrightarrow> d \<noteq> \<zero> \<Longrightarrow> (a \<oslash> b) \<otimes> (c \<oslash> d) = (a \<otimes> c) \<oslash> (b \<otimes> d)" | 
| 64962 | 384 | by (simp add: m_div_def nonzero_inverse_mult_distrib m_ac integral_iff) | 
| 385 | ||
| 386 | lemma nonzero_divide_divide_eq: | |
| 387 | "a \<in> carrier R \<Longrightarrow> b \<in> carrier R \<Longrightarrow> c \<in> carrier R \<Longrightarrow> d \<in> carrier R \<Longrightarrow> | |
| 67123 | 388 | b \<noteq> \<zero> \<Longrightarrow> c \<noteq> \<zero> \<Longrightarrow> d \<noteq> \<zero> \<Longrightarrow> (a \<oslash> b) \<oslash> (c \<oslash> d) = (a \<otimes> d) \<oslash> (b \<otimes> c)" | 
| 64962 | 389 | by (simp add: m_div_def nonzero_inverse_mult_distrib | 
| 390 | nonzero_imp_inverse_nonzero nonzero_inverse_inverse_eq m_ac integral_iff) | |
| 391 | ||
| 392 | lemma divide_1 [simp]: "x \<in> carrier R \<Longrightarrow> x \<oslash> \<one> = x" | |
| 393 | by (simp add: m_div_def) | |
| 394 | ||
| 395 | lemma add_frac_eq: | |
| 67123 | 396 | assumes "x \<in> carrier R" "y \<in> carrier R" "z \<in> carrier R" "w \<in> carrier R" | 
| 397 | and "y \<noteq> \<zero>" "z \<noteq> \<zero>" | |
| 64962 | 398 | shows "x \<oslash> y \<oplus> w \<oslash> z = (x \<otimes> z \<oplus> w \<otimes> y) \<oslash> (y \<otimes> z)" | 
| 399 | proof - | |
| 400 | from assms | |
| 401 | have "x \<oslash> y \<oplus> w \<oslash> z = x \<otimes> inv y \<otimes> (z \<otimes> inv z) \<oplus> w \<otimes> inv z \<otimes> (y \<otimes> inv y)" | |
| 402 | by (simp add: m_div_def) | |
| 403 | also from assms have "\<dots> = (x \<otimes> z \<oplus> w \<otimes> y) \<oslash> (y \<otimes> z)" | |
| 404 | by (simp add: m_div_def nonzero_inverse_mult_distrib r_distr m_ac integral_iff del: r_inv) | |
| 405 | finally show ?thesis . | |
| 406 | qed | |
| 407 | ||
| 408 | lemma div_closed [simp]: | |
| 409 | "x \<in> carrier R \<Longrightarrow> y \<in> carrier R \<Longrightarrow> y \<noteq> \<zero> \<Longrightarrow> x \<oslash> y \<in> carrier R" | |
| 410 | by (simp add: m_div_def) | |
| 411 | ||
| 412 | lemma minus_divide_left [simp]: | |
| 413 | "a \<in> carrier R \<Longrightarrow> b \<in> carrier R \<Longrightarrow> b \<noteq> \<zero> \<Longrightarrow> \<ominus> (a \<oslash> b) = \<ominus> a \<oslash> b" | |
| 414 | by (simp add: m_div_def l_minus) | |
| 415 | ||
| 416 | lemma diff_frac_eq: | |
| 67123 | 417 | assumes "x \<in> carrier R" "y \<in> carrier R" "z \<in> carrier R" "w \<in> carrier R" | 
| 418 | and "y \<noteq> \<zero>" "z \<noteq> \<zero>" | |
| 64962 | 419 | shows "x \<oslash> y \<ominus> w \<oslash> z = (x \<otimes> z \<ominus> w \<otimes> y) \<oslash> (y \<otimes> z)" | 
| 67123 | 420 | using assms by (simp add: minus_eq l_minus add_frac_eq) | 
| 64962 | 421 | |
| 422 | lemma nonzero_mult_divide_mult_cancel_left [simp]: | |
| 67123 | 423 | assumes "a \<in> carrier R" "b \<in> carrier R" "c \<in> carrier R" | 
| 424 | and "b \<noteq> \<zero>" "c \<noteq> \<zero>" | |
| 64962 | 425 | shows "(c \<otimes> a) \<oslash> (c \<otimes> b) = a \<oslash> b" | 
| 426 | proof - | |
| 427 | from assms have "(c \<otimes> a) \<oslash> (c \<otimes> b) = c \<otimes> a \<otimes> (inv b \<otimes> inv c)" | |
| 428 | by (simp add: m_div_def nonzero_inverse_mult_distrib integral_iff) | |
| 429 | also from assms have "\<dots> = a \<otimes> inv b \<otimes> (inv c \<otimes> c)" | |
| 430 | by (simp add: m_ac) | |
| 67123 | 431 | also from assms have "\<dots> = a \<otimes> inv b" | 
| 432 | by simp | |
| 433 | finally show ?thesis | |
| 434 | using assms by (simp add: m_div_def) | |
| 64962 | 435 | qed | 
| 436 | ||
| 437 | lemma times_divide_eq_left [simp]: | |
| 438 | "a \<in> carrier R \<Longrightarrow> b \<in> carrier R \<Longrightarrow> c \<in> carrier R \<Longrightarrow> c \<noteq> \<zero> \<Longrightarrow> | |
| 67123 | 439 | (b \<oslash> c) \<otimes> a = b \<otimes> a \<oslash> c" | 
| 64962 | 440 | by (simp add: m_div_def m_ac) | 
| 441 | ||
| 442 | lemma times_divide_eq_right [simp]: | |
| 443 | "a \<in> carrier R \<Longrightarrow> b \<in> carrier R \<Longrightarrow> c \<in> carrier R \<Longrightarrow> c \<noteq> \<zero> \<Longrightarrow> | |
| 67123 | 444 | a \<otimes> (b \<oslash> c) = a \<otimes> b \<oslash> c" | 
| 64962 | 445 | by (simp add: m_div_def m_ac) | 
| 446 | ||
| 447 | lemma nonzero_power_divide: | |
| 448 | "a \<in> carrier R \<Longrightarrow> b \<in> carrier R \<Longrightarrow> b \<noteq> \<zero> \<Longrightarrow> | |
| 67341 
df79ef3b3a41
Renamed (^) to [^] in preparation of the move from "op X" to (X)
 nipkow parents: 
67123diff
changeset | 449 | (a \<oslash> b) [^] (n::nat) = a [^] n \<oslash> b [^] n" | 
| 64962 | 450 | by (induct n) (simp_all add: nonzero_divide_divide_eq_left) | 
| 451 | ||
| 452 | lemma r_diff_distr: | |
| 453 | "x \<in> carrier R \<Longrightarrow> y \<in> carrier R \<Longrightarrow> z \<in> carrier R \<Longrightarrow> | |
| 67123 | 454 | z \<otimes> (x \<ominus> y) = z \<otimes> x \<ominus> z \<otimes> y" | 
| 64962 | 455 | by (simp add: minus_eq r_distr r_minus) | 
| 456 | ||
| 67123 | 457 | lemma divide_zero_left [simp]: "a \<in> carrier R \<Longrightarrow> a \<noteq> \<zero> \<Longrightarrow> \<zero> \<oslash> a = \<zero>" | 
| 64962 | 458 | by (simp add: m_div_def) | 
| 459 | ||
| 460 | lemma divide_self: "a \<in> carrier R \<Longrightarrow> a \<noteq> \<zero> \<Longrightarrow> a \<oslash> a = \<one>" | |
| 461 | by (simp add: m_div_def) | |
| 462 | ||
| 463 | lemma divide_eq_0_iff: | |
| 67123 | 464 | assumes "a \<in> carrier R" "b \<in> carrier R" | 
| 465 | and "b \<noteq> \<zero>" | |
| 466 | shows "a \<oslash> b = \<zero> \<longleftrightarrow> a = \<zero>" | |
| 64962 | 467 | proof | 
| 468 | assume "a = \<zero>" | |
| 469 | with assms show "a \<oslash> b = \<zero>" by simp | |
| 470 | next | |
| 471 | assume "a \<oslash> b = \<zero>" | |
| 472 | with assms have "b \<otimes> (a \<oslash> b) = b \<otimes> \<zero>" by simp | |
| 473 | also from assms have "b \<otimes> (a \<oslash> b) = b \<otimes> a \<oslash> b" by simp | |
| 474 | also from assms have "b \<otimes> a = a \<otimes> b" by (simp add: m_comm) | |
| 475 | also from assms have "a \<otimes> b \<oslash> b = a \<otimes> (b \<oslash> b)" by simp | |
| 476 | also from assms have "b \<oslash> b = \<one>" by (simp add: divide_self) | |
| 477 | finally show "a = \<zero>" using assms by simp | |
| 478 | qed | |
| 479 | ||
| 480 | end | |
| 481 | ||
| 482 | lemma field_class: "field (cring_class_ops::'a::field ring)" | |
| 483 | apply unfold_locales | |
| 67123 | 484 | apply (simp_all add: cring_class_ops_def) | 
| 64962 | 485 | apply (auto simp add: Units_def) | 
| 486 | apply (rule_tac x="1 / x" in exI) | |
| 487 | apply simp | |
| 488 | done | |
| 489 | ||
| 490 | lemma div_class: "(x::'a::field) \<oslash>\<^bsub>cring_class_ops\<^esub> y = x / y" | |
| 491 | apply (simp add: m_div_def m_inv_def class_simps) | |
| 492 | apply (rule impI) | |
| 493 | apply (rule ssubst [OF the_equality, of _ "1 / y"]) | |
| 67123 | 494 | apply simp_all | 
| 64962 | 495 | apply (drule conjunct2) | 
| 496 | apply (drule_tac f="\<lambda>x. x / y" in arg_cong) | |
| 497 | apply simp | |
| 498 | done | |
| 499 | ||
| 500 | interpretation field_class: field "cring_class_ops::'a::field ring" | |
| 67123 | 501 | rewrites "(\<zero>\<^bsub>cring_class_ops\<^esub>::'a) = 0" | 
| 502 | and "(\<one>\<^bsub>cring_class_ops\<^esub>::'a) = 1" | |
| 503 | and "(x::'a) \<oplus>\<^bsub>cring_class_ops\<^esub> y = x + y" | |
| 504 | and "(x::'a) \<otimes>\<^bsub>cring_class_ops\<^esub> y = x * y" | |
| 505 | and "\<ominus>\<^bsub>cring_class_ops\<^esub> (x::'a) = - x" | |
| 506 | and "(x::'a) \<ominus>\<^bsub>cring_class_ops\<^esub> y = x - y" | |
| 67341 
df79ef3b3a41
Renamed (^) to [^] in preparation of the move from "op X" to (X)
 nipkow parents: 
67123diff
changeset | 507 | and "(x::'a) [^]\<^bsub>cring_class_ops\<^esub> n = x ^ n" | 
| 67123 | 508 | and "\<guillemotleft>n\<guillemotright>\<^sub>\<nat>\<^bsub>cring_class_ops\<^esub> = of_nat n" | 
| 509 | and "((\<guillemotleft>i\<guillemotright>\<^bsub>cring_class_ops\<^esub>)::'a) = of_int i" | |
| 510 | and "(x::'a) \<oslash>\<^bsub>cring_class_ops\<^esub> y = x / y" | |
| 511 | and "(Int.of_int (numeral m)::'a) = numeral m" | |
| 64962 | 512 | by (simp_all add: field_class class_simps div_class) | 
| 513 | ||
| 514 | end |