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