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