| author | bulwahn | 
| Fri, 22 Jan 2010 15:26:29 +0100 | |
| changeset 34958 | dcd0fa5cc6d3 | 
| parent 33026 | 8f35633c4922 | 
| child 35317 | d57da4abb47d | 
| permissions | -rw-r--r-- | 
| 33026 | 1 | (* Title: HOL/Isar_Examples/Group.thy | 
| 6763 
df12aef00932
Some bits of group theory.  Demonstrate calculational proofs.
 wenzelm parents: diff
changeset | 2 | Author: Markus Wenzel, TU Muenchen | 
| 
df12aef00932
Some bits of group theory.  Demonstrate calculational proofs.
 wenzelm parents: diff
changeset | 3 | *) | 
| 
df12aef00932
Some bits of group theory.  Demonstrate calculational proofs.
 wenzelm parents: diff
changeset | 4 | |
| 10007 | 5 | header {* Basic group theory *}
 | 
| 7748 | 6 | |
| 31758 | 7 | theory Group | 
| 8 | imports Main | |
| 9 | begin | |
| 6763 
df12aef00932
Some bits of group theory.  Demonstrate calculational proofs.
 wenzelm parents: diff
changeset | 10 | |
| 10007 | 11 | subsection {* Groups and calculational reasoning *} 
 | 
| 6784 | 12 | |
| 13 | text {*
 | |
| 7968 | 14 |  Groups over signature $({\times} :: \alpha \To \alpha \To \alpha,
 | 
| 10141 | 15 |  \idt{one} :: \alpha, \idt{inverse} :: \alpha \To \alpha)$ are defined
 | 
| 16 | as an axiomatic type class as follows. Note that the parent class | |
| 7968 | 17 |  $\idt{times}$ is provided by the basic HOL theory.
 | 
| 10007 | 18 | *} | 
| 6763 
df12aef00932
Some bits of group theory.  Demonstrate calculational proofs.
 wenzelm parents: diff
changeset | 19 | |
| 
df12aef00932
Some bits of group theory.  Demonstrate calculational proofs.
 wenzelm parents: diff
changeset | 20 | consts | 
| 6784 | 21 | one :: "'a" | 
| 10141 | 22 | inverse :: "'a => 'a" | 
| 6763 
df12aef00932
Some bits of group theory.  Demonstrate calculational proofs.
 wenzelm parents: diff
changeset | 23 | |
| 
df12aef00932
Some bits of group theory.  Demonstrate calculational proofs.
 wenzelm parents: diff
changeset | 24 | axclass | 
| 
df12aef00932
Some bits of group theory.  Demonstrate calculational proofs.
 wenzelm parents: diff
changeset | 25 | group < times | 
| 6793 | 26 | group_assoc: "(x * y) * z = x * (y * z)" | 
| 10141 | 27 | group_left_one: "one * x = x" | 
| 28 | group_left_inverse: "inverse x * x = one" | |
| 6763 
df12aef00932
Some bits of group theory.  Demonstrate calculational proofs.
 wenzelm parents: diff
changeset | 29 | |
| 6784 | 30 | text {*
 | 
| 10141 | 31 | The group axioms only state the properties of left one and inverse, | 
| 7874 | 32 | the right versions may be derived as follows. | 
| 10007 | 33 | *} | 
| 6763 
df12aef00932
Some bits of group theory.  Demonstrate calculational proofs.
 wenzelm parents: diff
changeset | 34 | |
| 10141 | 35 | theorem group_right_inverse: "x * inverse x = (one::'a::group)" | 
| 10007 | 36 | proof - | 
| 10141 | 37 | have "x * inverse x = one * (x * inverse x)" | 
| 38 | by (simp only: group_left_one) | |
| 39 | also have "... = one * x * inverse x" | |
| 10007 | 40 | by (simp only: group_assoc) | 
| 10141 | 41 | also have "... = inverse (inverse x) * inverse x * x * inverse x" | 
| 10007 | 42 | by (simp only: group_left_inverse) | 
| 10141 | 43 | also have "... = inverse (inverse x) * (inverse x * x) * inverse x" | 
| 10007 | 44 | by (simp only: group_assoc) | 
| 10141 | 45 | also have "... = inverse (inverse x) * one * inverse x" | 
| 10007 | 46 | by (simp only: group_left_inverse) | 
| 10141 | 47 | also have "... = inverse (inverse x) * (one * inverse x)" | 
| 10007 | 48 | by (simp only: group_assoc) | 
| 10141 | 49 | also have "... = inverse (inverse x) * inverse x" | 
| 50 | by (simp only: group_left_one) | |
| 10007 | 51 | also have "... = one" | 
| 52 | by (simp only: group_left_inverse) | |
| 53 | finally show ?thesis . | |
| 54 | qed | |
| 6763 
df12aef00932
Some bits of group theory.  Demonstrate calculational proofs.
 wenzelm parents: diff
changeset | 55 | |
| 6784 | 56 | text {*
 | 
| 7982 | 57 |  With \name{group-right-inverse} already available,
 | 
| 10141 | 58 |  \name{group-right-one}\label{thm:group-right-one} is now established
 | 
| 59 | much easier. | |
| 10007 | 60 | *} | 
| 6763 
df12aef00932
Some bits of group theory.  Demonstrate calculational proofs.
 wenzelm parents: diff
changeset | 61 | |
| 10141 | 62 | theorem group_right_one: "x * one = (x::'a::group)" | 
| 10007 | 63 | proof - | 
| 10141 | 64 | have "x * one = x * (inverse x * x)" | 
| 10007 | 65 | by (simp only: group_left_inverse) | 
| 10141 | 66 | also have "... = x * inverse x * x" | 
| 10007 | 67 | by (simp only: group_assoc) | 
| 68 | also have "... = one * x" | |
| 69 | by (simp only: group_right_inverse) | |
| 70 | also have "... = x" | |
| 10141 | 71 | by (simp only: group_left_one) | 
| 10007 | 72 | finally show ?thesis . | 
| 73 | qed | |
| 6763 
df12aef00932
Some bits of group theory.  Demonstrate calculational proofs.
 wenzelm parents: diff
changeset | 74 | |
| 6784 | 75 | text {*
 | 
| 7874 | 76 | \medskip The calculational proof style above follows typical | 
| 77 | presentations given in any introductory course on algebra. The basic | |
| 78 | technique is to form a transitive chain of equations, which in turn | |
| 79 | are established by simplifying with appropriate rules. The low-level | |
| 7982 | 80 | logical details of equational reasoning are left implicit. | 
| 6784 | 81 | |
| 7982 | 82 | Note that ``$\dots$'' is just a special term variable that is bound | 
| 83 |  automatically to the argument\footnote{The argument of a curried
 | |
| 84 | infix expression happens to be its right-hand side.} of the last fact | |
| 85 | achieved by any local assumption or proven statement. In contrast to | |
| 86 |  $\var{thesis}$, the ``$\dots$'' variable is bound \emph{after} the
 | |
| 87 | proof is finished, though. | |
| 7874 | 88 | |
| 89 | There are only two separate Isar language elements for calculational | |
| 90 |  proofs: ``\isakeyword{also}'' for initial or intermediate
 | |
| 91 |  calculational steps, and ``\isakeyword{finally}'' for exhibiting the
 | |
| 92 | result of a calculation. These constructs are not hardwired into | |
| 93 | Isabelle/Isar, but defined on top of the basic Isar/VM interpreter. | |
| 94 |  Expanding the \isakeyword{also} and \isakeyword{finally} derived
 | |
| 7982 | 95 | language elements, calculations may be simulated by hand as | 
| 96 | demonstrated below. | |
| 10007 | 97 | *} | 
| 6784 | 98 | |
| 10007 | 99 | theorem "x * one = (x::'a::group)" | 
| 100 | proof - | |
| 10141 | 101 | have "x * one = x * (inverse x * x)" | 
| 10007 | 102 | by (simp only: group_left_inverse) | 
| 6763 
df12aef00932
Some bits of group theory.  Demonstrate calculational proofs.
 wenzelm parents: diff
changeset | 103 | |
| 7433 | 104 | note calculation = this | 
| 10007 | 105 |     -- {* first calculational step: init calculation register *}
 | 
| 6763 
df12aef00932
Some bits of group theory.  Demonstrate calculational proofs.
 wenzelm parents: diff
changeset | 106 | |
| 10141 | 107 | have "... = x * inverse x * x" | 
| 10007 | 108 | by (simp only: group_assoc) | 
| 6763 
df12aef00932
Some bits of group theory.  Demonstrate calculational proofs.
 wenzelm parents: diff
changeset | 109 | |
| 7433 | 110 | note calculation = trans [OF calculation this] | 
| 10007 | 111 |     -- {* general calculational step: compose with transitivity rule *}
 | 
| 6763 
df12aef00932
Some bits of group theory.  Demonstrate calculational proofs.
 wenzelm parents: diff
changeset | 112 | |
| 10007 | 113 | have "... = one * x" | 
| 114 | by (simp only: group_right_inverse) | |
| 6763 
df12aef00932
Some bits of group theory.  Demonstrate calculational proofs.
 wenzelm parents: diff
changeset | 115 | |
| 7433 | 116 | note calculation = trans [OF calculation this] | 
| 10007 | 117 |     -- {* general calculational step: compose with transitivity rule *}
 | 
| 6763 
df12aef00932
Some bits of group theory.  Demonstrate calculational proofs.
 wenzelm parents: diff
changeset | 118 | |
| 10007 | 119 | have "... = x" | 
| 10141 | 120 | by (simp only: group_left_one) | 
| 6763 
df12aef00932
Some bits of group theory.  Demonstrate calculational proofs.
 wenzelm parents: diff
changeset | 121 | |
| 7433 | 122 | note calculation = trans [OF calculation this] | 
| 10007 | 123 |     -- {* final calculational step: compose with transitivity rule ... *}
 | 
| 6784 | 124 | from calculation | 
| 10007 | 125 |     -- {* ... and pick up the final result *}
 | 
| 6763 
df12aef00932
Some bits of group theory.  Demonstrate calculational proofs.
 wenzelm parents: diff
changeset | 126 | |
| 10007 | 127 | show ?thesis . | 
| 128 | qed | |
| 6763 
df12aef00932
Some bits of group theory.  Demonstrate calculational proofs.
 wenzelm parents: diff
changeset | 129 | |
| 7874 | 130 | text {*
 | 
| 131 | Note that this scheme of calculations is not restricted to plain | |
| 132 | transitivity. Rules like anti-symmetry, or even forward and backward | |
| 7982 | 133 | substitution work as well. For the actual implementation of | 
| 134 |  \isacommand{also} and \isacommand{finally}, Isabelle/Isar maintains
 | |
| 135 | separate context information of ``transitivity'' rules. Rule | |
| 136 | selection takes place automatically by higher-order unification. | |
| 10007 | 137 | *} | 
| 7874 | 138 | |
| 6763 
df12aef00932
Some bits of group theory.  Demonstrate calculational proofs.
 wenzelm parents: diff
changeset | 139 | |
| 10007 | 140 | subsection {* Groups as monoids *}
 | 
| 6793 | 141 | |
| 142 | text {*
 | |
| 7874 | 143 |  Monoids over signature $({\times} :: \alpha \To \alpha \To \alpha,
 | 
| 144 |  \idt{one} :: \alpha)$ are defined like this.
 | |
| 10007 | 145 | *} | 
| 6793 | 146 | |
| 147 | axclass monoid < times | |
| 148 | monoid_assoc: "(x * y) * z = x * (y * z)" | |
| 10141 | 149 | monoid_left_one: "one * x = x" | 
| 150 | monoid_right_one: "x * one = x" | |
| 6793 | 151 | |
| 152 | text {*
 | |
| 7748 | 153 |  Groups are \emph{not} yet monoids directly from the definition.  For
 | 
| 10141 | 154 |  monoids, \name{right-one} had to be included as an axiom, but for
 | 
| 155 |  groups both \name{right-one} and \name{right-inverse} are derivable
 | |
| 156 |  from the other axioms.  With \name{group-right-one} derived as a
 | |
| 157 |  theorem of group theory (see page~\pageref{thm:group-right-one}), we
 | |
| 158 |  may still instantiate $\idt{group} \subseteq \idt{monoid}$ properly
 | |
| 159 | as follows. | |
| 10007 | 160 | *} | 
| 6793 | 161 | |
| 10007 | 162 | instance group < monoid | 
| 7356 | 163 | by (intro_classes, | 
| 6793 | 164 | rule group_assoc, | 
| 10141 | 165 | rule group_left_one, | 
| 166 | rule group_right_one) | |
| 6793 | 167 | |
| 7874 | 168 | text {*
 | 
| 169 |  The \isacommand{instance} command actually is a version of
 | |
| 170 |  \isacommand{theorem}, setting up a goal that reflects the intended
 | |
| 171 | class relation (or type constructor arity). Thus any Isar proof | |
| 172 | language element may be involved to establish this statement. When | |
| 7982 | 173 | concluding the proof, the result is transformed into the intended | 
| 7874 | 174 | type signature extension behind the scenes. | 
| 10007 | 175 | *} | 
| 7874 | 176 | |
| 10141 | 177 | subsection {* More theorems of group theory *}
 | 
| 178 | ||
| 179 | text {*
 | |
| 180 | The one element is already uniquely determined by preserving an | |
| 181 |  \emph{arbitrary} group element.
 | |
| 182 | *} | |
| 183 | ||
| 184 | theorem group_one_equality: "e * x = x ==> one = (e::'a::group)" | |
| 185 | proof - | |
| 186 | assume eq: "e * x = x" | |
| 187 | have "one = x * inverse x" | |
| 188 | by (simp only: group_right_inverse) | |
| 189 | also have "... = (e * x) * inverse x" | |
| 190 | by (simp only: eq) | |
| 191 | also have "... = e * (x * inverse x)" | |
| 192 | by (simp only: group_assoc) | |
| 193 | also have "... = e * one" | |
| 194 | by (simp only: group_right_inverse) | |
| 195 | also have "... = e" | |
| 196 | by (simp only: group_right_one) | |
| 197 | finally show ?thesis . | |
| 198 | qed | |
| 199 | ||
| 200 | text {*
 | |
| 201 | Likewise, the inverse is already determined by the cancel property. | |
| 202 | *} | |
| 203 | ||
| 204 | theorem group_inverse_equality: | |
| 205 | "x' * x = one ==> inverse x = (x'::'a::group)" | |
| 206 | proof - | |
| 207 | assume eq: "x' * x = one" | |
| 208 | have "inverse x = one * inverse x" | |
| 209 | by (simp only: group_left_one) | |
| 210 | also have "... = (x' * x) * inverse x" | |
| 211 | by (simp only: eq) | |
| 212 | also have "... = x' * (x * inverse x)" | |
| 213 | by (simp only: group_assoc) | |
| 214 | also have "... = x' * one" | |
| 215 | by (simp only: group_right_inverse) | |
| 216 | also have "... = x'" | |
| 217 | by (simp only: group_right_one) | |
| 218 | finally show ?thesis . | |
| 219 | qed | |
| 220 | ||
| 221 | text {*
 | |
| 222 | The inverse operation has some further characteristic properties. | |
| 223 | *} | |
| 224 | ||
| 225 | theorem group_inverse_times: | |
| 226 | "inverse (x * y) = inverse y * inverse (x::'a::group)" | |
| 227 | proof (rule group_inverse_equality) | |
| 228 | show "(inverse y * inverse x) * (x * y) = one" | |
| 229 | proof - | |
| 230 | have "(inverse y * inverse x) * (x * y) = | |
| 231 | (inverse y * (inverse x * x)) * y" | |
| 232 | by (simp only: group_assoc) | |
| 233 | also have "... = (inverse y * one) * y" | |
| 234 | by (simp only: group_left_inverse) | |
| 235 | also have "... = inverse y * y" | |
| 236 | by (simp only: group_right_one) | |
| 237 | also have "... = one" | |
| 238 | by (simp only: group_left_inverse) | |
| 239 | finally show ?thesis . | |
| 240 | qed | |
| 241 | qed | |
| 242 | ||
| 243 | theorem inverse_inverse: "inverse (inverse x) = (x::'a::group)" | |
| 244 | proof (rule group_inverse_equality) | |
| 245 | show "x * inverse x = one" | |
| 246 | by (simp only: group_right_inverse) | |
| 247 | qed | |
| 248 | ||
| 249 | theorem inverse_inject: "inverse x = inverse y ==> x = (y::'a::group)" | |
| 250 | proof - | |
| 251 | assume eq: "inverse x = inverse y" | |
| 252 | have "x = x * one" | |
| 253 | by (simp only: group_right_one) | |
| 254 | also have "... = x * (inverse y * y)" | |
| 255 | by (simp only: group_left_inverse) | |
| 256 | also have "... = x * (inverse x * y)" | |
| 257 | by (simp only: eq) | |
| 258 | also have "... = (x * inverse x) * y" | |
| 259 | by (simp only: group_assoc) | |
| 260 | also have "... = one * y" | |
| 261 | by (simp only: group_right_inverse) | |
| 262 | also have "... = y" | |
| 263 | by (simp only: group_left_one) | |
| 264 | finally show ?thesis . | |
| 265 | qed | |
| 266 | ||
| 267 | end |