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