| author | wenzelm | 
| Thu, 31 Jul 2014 21:29:31 +0200 | |
| changeset 57833 | 2c2bae3da1c2 | 
| parent 57514 | bdc2c6b40bf2 | 
| child 60112 | 3eab4acaa035 | 
| permissions | -rw-r--r-- | 
| 35849 | 1 | (* Title: HOL/Algebra/IntRing.thy | 
| 2 | Author: Stephan Hohe, TU Muenchen | |
| 41433 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 3 | Author: Clemens Ballarin | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 4 | *) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 5 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 6 | theory IntRing | 
| 55157 | 7 | imports QuotRing Lattice Int "~~/src/HOL/Number_Theory/Primes" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 8 | begin | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 9 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 10 | section {* The Ring of Integers *}
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 11 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 12 | subsection {* Some properties of @{typ int} *}
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 13 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 14 | lemma dvds_eq_abseq: | 
| 55991 | 15 | fixes k :: int | 
| 16 | shows "l dvd k \<and> k dvd l \<longleftrightarrow> abs l = abs k" | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 17 | apply rule | 
| 33657 | 18 | apply (simp add: zdvd_antisym_abs) | 
| 33676 
802f5e233e48
moved lemma from Algebra/IntRing to Ring_and_Field
 nipkow parents: 
33657diff
changeset | 19 | apply (simp add: dvd_if_abs_eq) | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 20 | done | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 21 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 22 | |
| 27717 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 ballarin parents: 
27713diff
changeset | 23 | subsection {* @{text "\<Z>"}: The Set of Integers as Algebraic Structure *}
 | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 24 | |
| 55991 | 25 | abbreviation int_ring :: "int ring" ("\<Z>")
 | 
| 26 | where "int_ring \<equiv> \<lparr>carrier = UNIV, mult = op *, one = 1, zero = 0, add = op +\<rparr>" | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 27 | |
| 55991 | 28 | lemma int_Zcarr [intro!, simp]: "k \<in> carrier \<Z>" | 
| 41433 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 29 | by simp | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 30 | |
| 55991 | 31 | lemma int_is_cring: "cring \<Z>" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 32 | apply (rule cringI) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 33 | apply (rule abelian_groupI, simp_all) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 34 | defer 1 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 35 | apply (rule comm_monoidI, simp_all) | 
| 49962 
a8cc904a6820
Renamed {left,right}_distrib to distrib_{right,left}.
 webertj parents: 
44821diff
changeset | 36 | apply (rule distrib_right) | 
| 44821 | 37 | apply (fast intro: left_minus) | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 38 | done | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 39 | |
| 23957 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 40 | (* | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 41 | lemma int_is_domain: | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 42 | "domain \<Z>" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 43 | apply (intro domain.intro domain_axioms.intro) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 44 | apply (rule int_is_cring) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 45 | apply (unfold int_ring_def, simp+) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 46 | done | 
| 23957 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 47 | *) | 
| 35849 | 48 | |
| 49 | ||
| 27717 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 ballarin parents: 
27713diff
changeset | 50 | subsection {* Interpretations *}
 | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 51 | |
| 23957 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 52 | text {* Since definitions of derived operations are global, their
 | 
| 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 53 | interpretation needs to be done as early as possible --- that is, | 
| 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 54 | with as few assumptions as possible. *} | 
| 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 55 | |
| 30729 
461ee3e49ad3
interpretation/interpret: prefixes are mandatory by default;
 wenzelm parents: 
29948diff
changeset | 56 | interpretation int: monoid \<Z> | 
| 23957 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 57 | where "carrier \<Z> = UNIV" | 
| 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 58 | and "mult \<Z> x y = x * y" | 
| 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 59 | and "one \<Z> = 1" | 
| 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 60 | and "pow \<Z> x n = x^n" | 
| 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 61 | proof - | 
| 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 62 | -- "Specification" | 
| 44655 | 63 | show "monoid \<Z>" by default auto | 
| 30729 
461ee3e49ad3
interpretation/interpret: prefixes are mandatory by default;
 wenzelm parents: 
29948diff
changeset | 64 | then interpret int: monoid \<Z> . | 
| 23957 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 65 | |
| 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 66 | -- "Carrier" | 
| 41433 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 67 | show "carrier \<Z> = UNIV" by simp | 
| 23957 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 68 | |
| 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 69 | -- "Operations" | 
| 41433 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 70 |   { fix x y show "mult \<Z> x y = x * y" by simp }
 | 
| 55991 | 71 | show "one \<Z> = 1" by simp | 
| 41433 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 72 | show "pow \<Z> x n = x^n" by (induct n) simp_all | 
| 23957 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 73 | qed | 
| 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 74 | |
| 30729 
461ee3e49ad3
interpretation/interpret: prefixes are mandatory by default;
 wenzelm parents: 
29948diff
changeset | 75 | interpretation int: comm_monoid \<Z> | 
| 28524 | 76 | where "finprod \<Z> f A = (if finite A then setprod f A else undefined)" | 
| 23957 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 77 | proof - | 
| 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 78 | -- "Specification" | 
| 44655 | 79 | show "comm_monoid \<Z>" by default auto | 
| 30729 
461ee3e49ad3
interpretation/interpret: prefixes are mandatory by default;
 wenzelm parents: 
29948diff
changeset | 80 | then interpret int: comm_monoid \<Z> . | 
| 23957 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 81 | |
| 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 82 | -- "Operations" | 
| 41433 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 83 |   { fix x y have "mult \<Z> x y = x * y" by simp }
 | 
| 23957 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 84 | note mult = this | 
| 41433 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 85 | have one: "one \<Z> = 1" by simp | 
| 28524 | 86 | show "finprod \<Z> f A = (if finite A then setprod f A else undefined)" | 
| 23957 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 87 | proof (cases "finite A") | 
| 55991 | 88 | case True | 
| 89 | then show ?thesis | |
| 90 | proof induct | |
| 91 | case empty | |
| 92 | show ?case by (simp add: one) | |
| 23957 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 93 | next | 
| 55991 | 94 | case insert | 
| 95 | then show ?case by (simp add: Pi_def mult) | |
| 23957 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 96 | qed | 
| 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 97 | next | 
| 55991 | 98 | case False | 
| 99 | then show ?thesis by (simp add: finprod_def) | |
| 23957 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 100 | qed | 
| 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 101 | qed | 
| 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 102 | |
| 30729 
461ee3e49ad3
interpretation/interpret: prefixes are mandatory by default;
 wenzelm parents: 
29948diff
changeset | 103 | interpretation int: abelian_monoid \<Z> | 
| 41433 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 104 | where int_carrier_eq: "carrier \<Z> = UNIV" | 
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 105 | and int_zero_eq: "zero \<Z> = 0" | 
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 106 | and int_add_eq: "add \<Z> x y = x + y" | 
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 107 | and int_finsum_eq: "finsum \<Z> f A = (if finite A then setsum f A else undefined)" | 
| 23957 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 108 | proof - | 
| 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 109 | -- "Specification" | 
| 44655 | 110 | show "abelian_monoid \<Z>" by default auto | 
| 30729 
461ee3e49ad3
interpretation/interpret: prefixes are mandatory by default;
 wenzelm parents: 
29948diff
changeset | 111 | then interpret int: abelian_monoid \<Z> . | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 112 | |
| 41433 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 113 | -- "Carrier" | 
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 114 | show "carrier \<Z> = UNIV" by simp | 
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 115 | |
| 23957 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 116 | -- "Operations" | 
| 41433 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 117 |   { fix x y show "add \<Z> x y = x + y" by simp }
 | 
| 23957 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 118 | note add = this | 
| 55991 | 119 | show zero: "zero \<Z> = 0" | 
| 120 | by simp | |
| 28524 | 121 | show "finsum \<Z> f A = (if finite A then setsum f A else undefined)" | 
| 23957 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 122 | proof (cases "finite A") | 
| 55991 | 123 | case True | 
| 124 | then show ?thesis | |
| 125 | proof induct | |
| 126 | case empty | |
| 127 | show ?case by (simp add: zero) | |
| 23957 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 128 | next | 
| 55991 | 129 | case insert | 
| 130 | then show ?case by (simp add: Pi_def add) | |
| 23957 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 131 | qed | 
| 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 132 | next | 
| 55991 | 133 | case False | 
| 134 | then show ?thesis by (simp add: finsum_def finprod_def) | |
| 23957 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 135 | qed | 
| 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 136 | qed | 
| 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 137 | |
| 30729 
461ee3e49ad3
interpretation/interpret: prefixes are mandatory by default;
 wenzelm parents: 
29948diff
changeset | 138 | interpretation int: abelian_group \<Z> | 
| 41433 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 139 | (* The equations from the interpretation of abelian_monoid need to be repeated. | 
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 140 | Since the morphisms through which the abelian structures are interpreted are | 
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 141 | not the identity, the equations of these interpretations are not inherited. *) | 
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 142 | (* FIXME *) | 
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 143 | where "carrier \<Z> = UNIV" | 
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 144 | and "zero \<Z> = 0" | 
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 145 | and "add \<Z> x y = x + y" | 
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 146 | and "finsum \<Z> f A = (if finite A then setsum f A else undefined)" | 
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 147 | and int_a_inv_eq: "a_inv \<Z> x = - x" | 
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 148 | and int_a_minus_eq: "a_minus \<Z> x y = x - y" | 
| 23957 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 149 | proof - | 
| 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 150 | -- "Specification" | 
| 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 151 | show "abelian_group \<Z>" | 
| 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 152 | proof (rule abelian_groupI) | 
| 55991 | 153 | fix x | 
| 154 | assume "x \<in> carrier \<Z>" | |
| 155 | then show "\<exists>y \<in> carrier \<Z>. y \<oplus>\<^bsub>\<Z>\<^esub> x = \<zero>\<^bsub>\<Z>\<^esub>" | |
| 41433 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 156 | by simp arith | 
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 157 | qed auto | 
| 30729 
461ee3e49ad3
interpretation/interpret: prefixes are mandatory by default;
 wenzelm parents: 
29948diff
changeset | 158 | then interpret int: abelian_group \<Z> . | 
| 23957 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 159 | -- "Operations" | 
| 41433 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 160 |   { fix x y have "add \<Z> x y = x + y" by simp }
 | 
| 23957 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 161 | note add = this | 
| 41433 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 162 | have zero: "zero \<Z> = 0" by simp | 
| 55991 | 163 |   {
 | 
| 164 | fix x | |
| 165 | have "add \<Z> (- x) x = zero \<Z>" | |
| 166 | by (simp add: add zero) | |
| 167 | then show "a_inv \<Z> x = - x" | |
| 168 | by (simp add: int.minus_equality) | |
| 169 | } | |
| 23957 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 170 | note a_inv = this | 
| 55991 | 171 | show "a_minus \<Z> x y = x - y" | 
| 172 | by (simp add: int.minus_eq add a_inv) | |
| 41433 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 173 | qed (simp add: int_carrier_eq int_zero_eq int_add_eq int_finsum_eq)+ | 
| 23957 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 174 | |
| 30729 
461ee3e49ad3
interpretation/interpret: prefixes are mandatory by default;
 wenzelm parents: 
29948diff
changeset | 175 | interpretation int: "domain" \<Z> | 
| 41433 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 176 | where "carrier \<Z> = UNIV" | 
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 177 | and "zero \<Z> = 0" | 
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 178 | and "add \<Z> x y = x + y" | 
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 179 | and "finsum \<Z> f A = (if finite A then setsum f A else undefined)" | 
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 180 | and "a_inv \<Z> x = - x" | 
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 181 | and "a_minus \<Z> x y = x - y" | 
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 182 | proof - | 
| 55991 | 183 | show "domain \<Z>" | 
| 184 | by unfold_locales (auto simp: distrib_right distrib_left) | |
| 185 | qed (simp add: int_carrier_eq int_zero_eq int_add_eq int_finsum_eq int_a_inv_eq int_a_minus_eq)+ | |
| 23957 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 186 | |
| 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 187 | |
| 24131 
1099f6c73649
Experimental removal of assumptions of the form x : UNIV and the like after interpretation.
 ballarin parents: 
23957diff
changeset | 188 | text {* Removal of occurrences of @{term UNIV} in interpretation result
 | 
| 
1099f6c73649
Experimental removal of assumptions of the form x : UNIV and the like after interpretation.
 ballarin parents: 
23957diff
changeset | 189 | --- experimental. *} | 
| 
1099f6c73649
Experimental removal of assumptions of the form x : UNIV and the like after interpretation.
 ballarin parents: 
23957diff
changeset | 190 | |
| 
1099f6c73649
Experimental removal of assumptions of the form x : UNIV and the like after interpretation.
 ballarin parents: 
23957diff
changeset | 191 | lemma UNIV: | 
| 55991 | 192 | "x \<in> UNIV \<longleftrightarrow> True" | 
| 193 | "A \<subseteq> UNIV \<longleftrightarrow> True" | |
| 194 | "(\<forall>x \<in> UNIV. P x) \<longleftrightarrow> (\<forall>x. P x)" | |
| 195 | "(EX x : UNIV. P x) \<longleftrightarrow> (EX x. P x)" | |
| 196 | "(True \<longrightarrow> Q) \<longleftrightarrow> Q" | |
| 197 | "(True \<Longrightarrow> PROP R) \<equiv> PROP R" | |
| 24131 
1099f6c73649
Experimental removal of assumptions of the form x : UNIV and the like after interpretation.
 ballarin parents: 
23957diff
changeset | 198 | by simp_all | 
| 
1099f6c73649
Experimental removal of assumptions of the form x : UNIV and the like after interpretation.
 ballarin parents: 
23957diff
changeset | 199 | |
| 30729 
461ee3e49ad3
interpretation/interpret: prefixes are mandatory by default;
 wenzelm parents: 
29948diff
changeset | 200 | interpretation int (* FIXME [unfolded UNIV] *) : | 
| 55926 | 201 | partial_order "\<lparr>carrier = UNIV::int set, eq = op =, le = op \<le>\<rparr>" | 
| 202 | where "carrier \<lparr>carrier = UNIV::int set, eq = op =, le = op \<le>\<rparr> = UNIV" | |
| 203 | and "le \<lparr>carrier = UNIV::int set, eq = op =, le = op \<le>\<rparr> x y = (x \<le> y)" | |
| 204 | and "lless \<lparr>carrier = UNIV::int set, eq = op =, le = op \<le>\<rparr> x y = (x < y)" | |
| 23957 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 205 | proof - | 
| 55926 | 206 | show "partial_order \<lparr>carrier = UNIV::int set, eq = op =, le = op \<le>\<rparr>" | 
| 44655 | 207 | by default simp_all | 
| 55926 | 208 | show "carrier \<lparr>carrier = UNIV::int set, eq = op =, le = op \<le>\<rparr> = UNIV" | 
| 24131 
1099f6c73649
Experimental removal of assumptions of the form x : UNIV and the like after interpretation.
 ballarin parents: 
23957diff
changeset | 209 | by simp | 
| 55926 | 210 | show "le \<lparr>carrier = UNIV::int set, eq = op =, le = op \<le>\<rparr> x y = (x \<le> y)" | 
| 24131 
1099f6c73649
Experimental removal of assumptions of the form x : UNIV and the like after interpretation.
 ballarin parents: 
23957diff
changeset | 211 | by simp | 
| 55926 | 212 | show "lless \<lparr>carrier = UNIV::int set, eq = op =, le = op \<le>\<rparr> x y = (x < y)" | 
| 23957 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 213 | by (simp add: lless_def) auto | 
| 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 214 | qed | 
| 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 215 | |
| 30729 
461ee3e49ad3
interpretation/interpret: prefixes are mandatory by default;
 wenzelm parents: 
29948diff
changeset | 216 | interpretation int (* FIXME [unfolded UNIV] *) : | 
| 55926 | 217 | lattice "\<lparr>carrier = UNIV::int set, eq = op =, le = op \<le>\<rparr>" | 
| 218 | where "join \<lparr>carrier = UNIV::int set, eq = op =, le = op \<le>\<rparr> x y = max x y" | |
| 219 | and "meet \<lparr>carrier = UNIV::int set, eq = op =, le = op \<le>\<rparr> x y = min x y" | |
| 23957 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 220 | proof - | 
| 55926 | 221 | let ?Z = "\<lparr>carrier = UNIV::int set, eq = op =, le = op \<le>\<rparr>" | 
| 23957 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 222 | show "lattice ?Z" | 
| 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 223 | apply unfold_locales | 
| 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 224 | apply (simp add: least_def Upper_def) | 
| 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 225 | apply arith | 
| 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 226 | apply (simp add: greatest_def Lower_def) | 
| 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 227 | apply arith | 
| 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 228 | done | 
| 30729 
461ee3e49ad3
interpretation/interpret: prefixes are mandatory by default;
 wenzelm parents: 
29948diff
changeset | 229 | then interpret int: lattice "?Z" . | 
| 23957 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 230 | show "join ?Z x y = max x y" | 
| 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 231 | apply (rule int.joinI) | 
| 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 232 | apply (simp_all add: least_def Upper_def) | 
| 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 233 | apply arith | 
| 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 234 | done | 
| 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 235 | show "meet ?Z x y = min x y" | 
| 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 236 | apply (rule int.meetI) | 
| 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 237 | apply (simp_all add: greatest_def Lower_def) | 
| 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 238 | apply arith | 
| 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 239 | done | 
| 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 240 | qed | 
| 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 241 | |
| 30729 
461ee3e49ad3
interpretation/interpret: prefixes are mandatory by default;
 wenzelm parents: 
29948diff
changeset | 242 | interpretation int (* [unfolded UNIV] *) : | 
| 55926 | 243 | total_order "\<lparr>carrier = UNIV::int set, eq = op =, le = op \<le>\<rparr>" | 
| 44655 | 244 | by default clarsimp | 
| 23957 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 245 | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 246 | |
| 27717 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 ballarin parents: 
27713diff
changeset | 247 | subsection {* Generated Ideals of @{text "\<Z>"} *}
 | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 248 | |
| 55991 | 249 | lemma int_Idl: "Idl\<^bsub>\<Z>\<^esub> {a} = {x * a | x. True}"
 | 
| 41433 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 250 | apply (subst int.cgenideal_eq_genideal[symmetric]) apply simp | 
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 251 | apply (simp add: cgenideal_def) | 
| 23957 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 252 | done | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 253 | |
| 55991 | 254 | lemma multiples_principalideal: "principalideal {x * a | x. True } \<Z>"
 | 
| 255 | by (metis UNIV_I int.cgenideal_eq_genideal int.cgenideal_is_principalideal int_Idl) | |
| 29700 | 256 | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 257 | lemma prime_primeideal: | 
| 55242 
413ec965f95d
Number_Theory: prime is no longer overloaded, but only for nat. Automatic coercion to int enabled.
 paulson <lp15@cam.ac.uk> parents: 
55157diff
changeset | 258 | assumes prime: "prime p" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 259 |   shows "primeideal (Idl\<^bsub>\<Z>\<^esub> {p}) \<Z>"
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 260 | apply (rule primeidealI) | 
| 23957 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 261 | apply (rule int.genideal_ideal, simp) | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 262 | apply (rule int_is_cring) | 
| 23957 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 263 | apply (simp add: int.cgenideal_eq_genideal[symmetric] cgenideal_def) | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 264 | apply clarsimp defer 1 | 
| 23957 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 ballarin parents: 
22063diff
changeset | 265 | apply (simp add: int.cgenideal_eq_genideal[symmetric] cgenideal_def) | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 266 | apply (elim exE) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 267 | proof - | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 268 | fix a b x | 
| 55242 
413ec965f95d
Number_Theory: prime is no longer overloaded, but only for nat. Automatic coercion to int enabled.
 paulson <lp15@cam.ac.uk> parents: 
55157diff
changeset | 269 | assume "a * b = x * int p" | 
| 55991 | 270 | then have "p dvd a * b" by simp | 
| 271 | then have "p dvd a \<or> p dvd b" | |
| 55242 
413ec965f95d
Number_Theory: prime is no longer overloaded, but only for nat. Automatic coercion to int enabled.
 paulson <lp15@cam.ac.uk> parents: 
55157diff
changeset | 272 | by (metis prime prime_dvd_mult_eq_int) | 
| 55991 | 273 | then show "(\<exists>x. a = x * int p) \<or> (\<exists>x. b = x * int p)" | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
55991diff
changeset | 274 | by (metis dvd_def mult.commute) | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 275 | next | 
| 55242 
413ec965f95d
Number_Theory: prime is no longer overloaded, but only for nat. Automatic coercion to int enabled.
 paulson <lp15@cam.ac.uk> parents: 
55157diff
changeset | 276 |   assume "UNIV = {uu. EX x. uu = x * int p}"
 | 
| 
413ec965f95d
Number_Theory: prime is no longer overloaded, but only for nat. Automatic coercion to int enabled.
 paulson <lp15@cam.ac.uk> parents: 
55157diff
changeset | 277 | then obtain x where "1 = x * int p" by best | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
55991diff
changeset | 278 | then have "\<bar>int p * x\<bar> = 1" by (simp add: mult.commute) | 
| 55242 
413ec965f95d
Number_Theory: prime is no longer overloaded, but only for nat. Automatic coercion to int enabled.
 paulson <lp15@cam.ac.uk> parents: 
55157diff
changeset | 279 | then show False | 
| 
413ec965f95d
Number_Theory: prime is no longer overloaded, but only for nat. Automatic coercion to int enabled.
 paulson <lp15@cam.ac.uk> parents: 
55157diff
changeset | 280 | by (metis abs_of_nat int_1 of_nat_eq_iff abs_zmult_eq_1 one_not_prime_nat prime) | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 281 | qed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 282 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 283 | |
| 27717 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 ballarin parents: 
27713diff
changeset | 284 | subsection {* Ideals and Divisibility *}
 | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 285 | |
| 55991 | 286 | lemma int_Idl_subset_ideal: "Idl\<^bsub>\<Z>\<^esub> {k} \<subseteq> Idl\<^bsub>\<Z>\<^esub> {l} = (k \<in> Idl\<^bsub>\<Z>\<^esub> {l})"
 | 
| 287 | by (rule int.Idl_subset_ideal') simp_all | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 288 | |
| 55991 | 289 | lemma Idl_subset_eq_dvd: "Idl\<^bsub>\<Z>\<^esub> {k} \<subseteq> Idl\<^bsub>\<Z>\<^esub> {l} \<longleftrightarrow> l dvd k"
 | 
| 290 | apply (subst int_Idl_subset_ideal, subst int_Idl, simp) | |
| 291 | apply (rule, clarify) | |
| 292 | apply (simp add: dvd_def) | |
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 293 | apply (simp add: dvd_def ac_simps) | 
| 55991 | 294 | done | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 295 | |
| 55991 | 296 | lemma dvds_eq_Idl: "l dvd k \<and> k dvd l \<longleftrightarrow> Idl\<^bsub>\<Z>\<^esub> {k} = Idl\<^bsub>\<Z>\<^esub> {l}"
 | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 297 | proof - | 
| 55991 | 298 |   have a: "l dvd k \<longleftrightarrow> (Idl\<^bsub>\<Z>\<^esub> {k} \<subseteq> Idl\<^bsub>\<Z>\<^esub> {l})"
 | 
| 299 | by (rule Idl_subset_eq_dvd[symmetric]) | |
| 300 |   have b: "k dvd l \<longleftrightarrow> (Idl\<^bsub>\<Z>\<^esub> {l} \<subseteq> Idl\<^bsub>\<Z>\<^esub> {k})"
 | |
| 301 | by (rule Idl_subset_eq_dvd[symmetric]) | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 302 | |
| 55991 | 303 |   have "l dvd k \<and> k dvd l \<longleftrightarrow> Idl\<^bsub>\<Z>\<^esub> {k} \<subseteq> Idl\<^bsub>\<Z>\<^esub> {l} \<and> Idl\<^bsub>\<Z>\<^esub> {l} \<subseteq> Idl\<^bsub>\<Z>\<^esub> {k}"
 | 
| 304 | by (subst a, subst b, simp) | |
| 305 |   also have "Idl\<^bsub>\<Z>\<^esub> {k} \<subseteq> Idl\<^bsub>\<Z>\<^esub> {l} \<and> Idl\<^bsub>\<Z>\<^esub> {l} \<subseteq> Idl\<^bsub>\<Z>\<^esub> {k} \<longleftrightarrow> Idl\<^bsub>\<Z>\<^esub> {k} = Idl\<^bsub>\<Z>\<^esub> {l}"
 | |
| 306 | by blast | |
| 307 | finally show ?thesis . | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 308 | qed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 309 | |
| 55991 | 310 | lemma Idl_eq_abs: "Idl\<^bsub>\<Z>\<^esub> {k} = Idl\<^bsub>\<Z>\<^esub> {l} \<longleftrightarrow> abs l = abs k"
 | 
| 311 | apply (subst dvds_eq_abseq[symmetric]) | |
| 312 | apply (rule dvds_eq_Idl[symmetric]) | |
| 313 | done | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 314 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 315 | |
| 27717 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 ballarin parents: 
27713diff
changeset | 316 | subsection {* Ideals and the Modulus *}
 | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 317 | |
| 55991 | 318 | definition ZMod :: "int \<Rightarrow> int \<Rightarrow> int set" | 
| 35848 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 wenzelm parents: 
35416diff
changeset | 319 |   where "ZMod k r = (Idl\<^bsub>\<Z>\<^esub> {k}) +>\<^bsub>\<Z>\<^esub> r"
 | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 320 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 321 | lemmas ZMod_defs = | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 322 | ZMod_def genideal_def | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 323 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 324 | lemma rcos_zfact: | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 325 | assumes kIl: "k \<in> ZMod l r" | 
| 55991 | 326 | shows "\<exists>x. k = x * l + r" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 327 | proof - | 
| 55991 | 328 |   from kIl[unfolded ZMod_def] have "\<exists>xl\<in>Idl\<^bsub>\<Z>\<^esub> {l}. k = xl + r"
 | 
| 329 | by (simp add: a_r_coset_defs) | |
| 330 |   then obtain xl where xl: "xl \<in> Idl\<^bsub>\<Z>\<^esub> {l}" and k: "k = xl + r"
 | |
| 331 | by auto | |
| 332 | from xl obtain x where "xl = x * l" | |
| 333 | by (auto simp: int_Idl) | |
| 334 | with k have "k = x * l + r" | |
| 335 | by simp | |
| 336 | then show "\<exists>x. k = x * l + r" .. | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 337 | qed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 338 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 339 | lemma ZMod_imp_zmod: | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 340 | assumes zmods: "ZMod m a = ZMod m b" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 341 | shows "a mod m = b mod m" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 342 | proof - | 
| 55991 | 343 |   interpret ideal "Idl\<^bsub>\<Z>\<^esub> {m}" \<Z>
 | 
| 344 | by (rule int.genideal_ideal) fast | |
| 345 | from zmods have "b \<in> ZMod m a" | |
| 346 | unfolding ZMod_def by (simp add: a_repr_independenceD) | |
| 347 | then have "\<exists>x. b = x * m + a" | |
| 348 | by (rule rcos_zfact) | |
| 349 | then obtain x where "b = x * m + a" | |
| 350 | by fast | |
| 351 | then have "b mod m = (x * m + a) mod m" | |
| 352 | by simp | |
| 353 | also have "\<dots> = ((x * m) mod m) + (a mod m)" | |
| 354 | by (simp add: mod_add_eq) | |
| 355 | also have "\<dots> = a mod m" | |
| 356 | by simp | |
| 357 | finally have "b mod m = a mod m" . | |
| 358 | then show "a mod m = b mod m" .. | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 359 | qed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 360 | |
| 55991 | 361 | lemma ZMod_mod: "ZMod m a = ZMod m (a mod m)" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 362 | proof - | 
| 55991 | 363 |   interpret ideal "Idl\<^bsub>\<Z>\<^esub> {m}" \<Z>
 | 
| 364 | by (rule int.genideal_ideal) fast | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 365 | show ?thesis | 
| 55991 | 366 | unfolding ZMod_def | 
| 367 | apply (rule a_repr_independence'[symmetric]) | |
| 368 | apply (simp add: int_Idl a_r_coset_defs) | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 369 | proof - | 
| 55991 | 370 | have "a = m * (a div m) + (a mod m)" | 
| 371 | by (simp add: zmod_zdiv_equality) | |
| 372 | then have "a = (a div m) * m + (a mod m)" | |
| 373 | by simp | |
| 374 | then show "\<exists>h. (\<exists>x. h = x * m) \<and> a = h + a mod m" | |
| 375 | by fast | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 376 | qed simp | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 377 | qed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 378 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 379 | lemma zmod_imp_ZMod: | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 380 | assumes modeq: "a mod m = b mod m" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 381 | shows "ZMod m a = ZMod m b" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 382 | proof - | 
| 55991 | 383 | have "ZMod m a = ZMod m (a mod m)" | 
| 384 | by (rule ZMod_mod) | |
| 385 | also have "\<dots> = ZMod m (b mod m)" | |
| 386 | by (simp add: modeq[symmetric]) | |
| 387 | also have "\<dots> = ZMod m b" | |
| 388 | by (rule ZMod_mod[symmetric]) | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 389 | finally show ?thesis . | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 390 | qed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 391 | |
| 55991 | 392 | corollary ZMod_eq_mod: "ZMod m a = ZMod m b \<longleftrightarrow> a mod m = b mod m" | 
| 393 | apply (rule iffI) | |
| 394 | apply (erule ZMod_imp_zmod) | |
| 395 | apply (erule zmod_imp_ZMod) | |
| 396 | done | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 397 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 398 | |
| 27717 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 ballarin parents: 
27713diff
changeset | 399 | subsection {* Factorization *}
 | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 400 | |
| 55991 | 401 | definition ZFact :: "int \<Rightarrow> int set ring" | 
| 35848 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 wenzelm parents: 
35416diff
changeset | 402 |   where "ZFact k = \<Z> Quot (Idl\<^bsub>\<Z>\<^esub> {k})"
 | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 403 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 404 | lemmas ZFact_defs = ZFact_def FactRing_def | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 405 | |
| 55991 | 406 | lemma ZFact_is_cring: "cring (ZFact k)" | 
| 407 | apply (unfold ZFact_def) | |
| 408 | apply (rule ideal.quotient_is_cring) | |
| 409 | apply (intro ring.genideal_ideal) | |
| 410 | apply (simp add: cring.axioms[OF int_is_cring] ring.intro) | |
| 411 | apply simp | |
| 412 | apply (rule int_is_cring) | |
| 413 | done | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 414 | |
| 55991 | 415 | lemma ZFact_zero: "carrier (ZFact 0) = (\<Union>a. {{a}})"
 | 
| 416 | apply (insert int.genideal_zero) | |
| 417 | apply (simp add: ZFact_defs A_RCOSETS_defs r_coset_def) | |
| 418 | done | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 419 | |
| 55991 | 420 | lemma ZFact_one: "carrier (ZFact 1) = {UNIV}"
 | 
| 421 | apply (simp only: ZFact_defs A_RCOSETS_defs r_coset_def ring_record_simps) | |
| 422 | apply (subst int.genideal_one) | |
| 423 | apply (rule, rule, clarsimp) | |
| 424 | apply (rule, rule, clarsimp) | |
| 425 | apply (rule, clarsimp, arith) | |
| 426 | apply (rule, clarsimp) | |
| 427 | apply (rule exI[of _ "0"], clarsimp) | |
| 428 | done | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 429 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 430 | lemma ZFact_prime_is_domain: | 
| 55242 
413ec965f95d
Number_Theory: prime is no longer overloaded, but only for nat. Automatic coercion to int enabled.
 paulson <lp15@cam.ac.uk> parents: 
55157diff
changeset | 431 | assumes pprime: "prime p" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 432 | shows "domain (ZFact p)" | 
| 55991 | 433 | apply (unfold ZFact_def) | 
| 434 | apply (rule primeideal.quotient_is_domain) | |
| 435 | apply (rule prime_primeideal[OF pprime]) | |
| 436 | done | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 437 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 438 | end |