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