src/HOL/Isar_examples/Group.thy
author wenzelm
Wed, 06 Oct 1999 18:50:51 +0200
changeset 7761 7fab9592384f
parent 7748 5b9c45b21782
child 7800 8ee919e42174
permissions -rw-r--r--
improved presentation;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
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
7748
5b9c45b21782 improved presentation;
wenzelm
parents: 7740
diff changeset
     6
header {* Basic group theory *};
5b9c45b21782 improved presentation;
wenzelm
parents: 7740
diff changeset
     7
6763
df12aef00932 Some bits of group theory. Demonstrate calculational proofs.
wenzelm
parents:
diff changeset
     8
theory Group = Main:;
df12aef00932 Some bits of group theory. Demonstrate calculational proofs.
wenzelm
parents:
diff changeset
     9
7748
5b9c45b21782 improved presentation;
wenzelm
parents: 7740
diff changeset
    10
subsection {* Groups *}; 
6784
687ddcad8581 proper calculation;
wenzelm
parents: 6763
diff changeset
    11
687ddcad8581 proper calculation;
wenzelm
parents: 6763
diff changeset
    12
text {*
6793
88aba7f486cb groups as monoids;
wenzelm
parents: 6784
diff changeset
    13
 We define an axiomatic type class of general groups over signature
7748
5b9c45b21782 improved presentation;
wenzelm
parents: 7740
diff changeset
    14
 $({\times}, \idt{one}, \idt{inv})$.
6784
687ddcad8581 proper calculation;
wenzelm
parents: 6763
diff changeset
    15
*};
6763
df12aef00932 Some bits of group theory. Demonstrate calculational proofs.
wenzelm
parents:
diff changeset
    16
df12aef00932 Some bits of group theory. Demonstrate calculational proofs.
wenzelm
parents:
diff changeset
    17
consts
6784
687ddcad8581 proper calculation;
wenzelm
parents: 6763
diff changeset
    18
  one :: "'a"
687ddcad8581 proper calculation;
wenzelm
parents: 6763
diff changeset
    19
  inv :: "'a => 'a";
6763
df12aef00932 Some bits of group theory. Demonstrate calculational proofs.
wenzelm
parents:
diff changeset
    20
df12aef00932 Some bits of group theory. Demonstrate calculational proofs.
wenzelm
parents:
diff changeset
    21
axclass
df12aef00932 Some bits of group theory. Demonstrate calculational proofs.
wenzelm
parents:
diff changeset
    22
  group < times
6793
88aba7f486cb groups as monoids;
wenzelm
parents: 6784
diff changeset
    23
  group_assoc:         "(x * y) * z = x * (y * z)"
88aba7f486cb groups as monoids;
wenzelm
parents: 6784
diff changeset
    24
  group_left_unit:     "one * x = x"
88aba7f486cb groups as monoids;
wenzelm
parents: 6784
diff changeset
    25
  group_left_inverse:  "inv x * x = one";
6763
df12aef00932 Some bits of group theory. Demonstrate calculational proofs.
wenzelm
parents:
diff changeset
    26
6784
687ddcad8581 proper calculation;
wenzelm
parents: 6763
diff changeset
    27
text {*
6793
88aba7f486cb groups as monoids;
wenzelm
parents: 6784
diff changeset
    28
 The group axioms only state the properties of left unit and inverse,
88aba7f486cb groups as monoids;
wenzelm
parents: 6784
diff changeset
    29
 the right versions are derivable as follows.  The calculational proof
88aba7f486cb groups as monoids;
wenzelm
parents: 6784
diff changeset
    30
 style below closely follows typical presentations given in any basic
88aba7f486cb groups as monoids;
wenzelm
parents: 6784
diff changeset
    31
 course on algebra.
6784
687ddcad8581 proper calculation;
wenzelm
parents: 6763
diff changeset
    32
*};
6763
df12aef00932 Some bits of group theory. Demonstrate calculational proofs.
wenzelm
parents:
diff changeset
    33
6793
88aba7f486cb groups as monoids;
wenzelm
parents: 6784
diff changeset
    34
theorem group_right_inverse: "x * inv x = (one::'a::group)";
7133
64c9f2364dae renamed 'same' to '-';
wenzelm
parents: 7005
diff changeset
    35
proof -;
6784
687ddcad8581 proper calculation;
wenzelm
parents: 6763
diff changeset
    36
  have "x * inv x = one * (x * inv x)";
6908
1bf0590f4790 simp only;
wenzelm
parents: 6892
diff changeset
    37
    by (simp only: group_left_unit);
6784
687ddcad8581 proper calculation;
wenzelm
parents: 6763
diff changeset
    38
  also; have "... = (one * x) * inv x";
6908
1bf0590f4790 simp only;
wenzelm
parents: 6892
diff changeset
    39
    by (simp only: group_assoc);
6784
687ddcad8581 proper calculation;
wenzelm
parents: 6763
diff changeset
    40
  also; have "... = inv (inv x) * inv x * x * inv x";
6908
1bf0590f4790 simp only;
wenzelm
parents: 6892
diff changeset
    41
    by (simp only: group_left_inverse);
6784
687ddcad8581 proper calculation;
wenzelm
parents: 6763
diff changeset
    42
  also; have "... = inv (inv x) * (inv x * x) * inv x";
6908
1bf0590f4790 simp only;
wenzelm
parents: 6892
diff changeset
    43
    by (simp only: group_assoc);
6784
687ddcad8581 proper calculation;
wenzelm
parents: 6763
diff changeset
    44
  also; have "... = inv (inv x) * one * inv x";
6908
1bf0590f4790 simp only;
wenzelm
parents: 6892
diff changeset
    45
    by (simp only: group_left_inverse);
6784
687ddcad8581 proper calculation;
wenzelm
parents: 6763
diff changeset
    46
  also; have "... = inv (inv x) * (one * inv x)";
6908
1bf0590f4790 simp only;
wenzelm
parents: 6892
diff changeset
    47
    by (simp only: group_assoc);
6784
687ddcad8581 proper calculation;
wenzelm
parents: 6763
diff changeset
    48
  also; have "... = inv (inv x) * inv x";
6908
1bf0590f4790 simp only;
wenzelm
parents: 6892
diff changeset
    49
    by (simp only: group_left_unit);
6784
687ddcad8581 proper calculation;
wenzelm
parents: 6763
diff changeset
    50
  also; have "... = one";
6908
1bf0590f4790 simp only;
wenzelm
parents: 6892
diff changeset
    51
    by (simp only: group_left_inverse);
7480
0a0e0dbe1269 replaced ?? by ?;
wenzelm
parents: 7433
diff changeset
    52
  finally; show ?thesis; .;
6784
687ddcad8581 proper calculation;
wenzelm
parents: 6763
diff changeset
    53
qed;
6763
df12aef00932 Some bits of group theory. Demonstrate calculational proofs.
wenzelm
parents:
diff changeset
    54
6784
687ddcad8581 proper calculation;
wenzelm
parents: 6763
diff changeset
    55
text {*
7748
5b9c45b21782 improved presentation;
wenzelm
parents: 7740
diff changeset
    56
 With \name{group-right-inverse} already at our disposal,
5b9c45b21782 improved presentation;
wenzelm
parents: 7740
diff changeset
    57
 \name{group-right-unit} is now obtained much easier as follows.
6784
687ddcad8581 proper calculation;
wenzelm
parents: 6763
diff changeset
    58
*};
6763
df12aef00932 Some bits of group theory. Demonstrate calculational proofs.
wenzelm
parents:
diff changeset
    59
6793
88aba7f486cb groups as monoids;
wenzelm
parents: 6784
diff changeset
    60
theorem group_right_unit: "x * one = (x::'a::group)";
7133
64c9f2364dae renamed 'same' to '-';
wenzelm
parents: 7005
diff changeset
    61
proof -;
6784
687ddcad8581 proper calculation;
wenzelm
parents: 6763
diff changeset
    62
  have "x * one = x * (inv x * x)";
6908
1bf0590f4790 simp only;
wenzelm
parents: 6892
diff changeset
    63
    by (simp only: group_left_inverse);
6784
687ddcad8581 proper calculation;
wenzelm
parents: 6763
diff changeset
    64
  also; have "... = x * inv x * x";
6908
1bf0590f4790 simp only;
wenzelm
parents: 6892
diff changeset
    65
    by (simp only: group_assoc);
6784
687ddcad8581 proper calculation;
wenzelm
parents: 6763
diff changeset
    66
  also; have "... = one * x";
6908
1bf0590f4790 simp only;
wenzelm
parents: 6892
diff changeset
    67
    by (simp only: group_right_inverse);
6784
687ddcad8581 proper calculation;
wenzelm
parents: 6763
diff changeset
    68
  also; have "... = x";
6908
1bf0590f4790 simp only;
wenzelm
parents: 6892
diff changeset
    69
    by (simp only: group_left_unit);
7480
0a0e0dbe1269 replaced ?? by ?;
wenzelm
parents: 7433
diff changeset
    70
  finally; show ?thesis; .;
6784
687ddcad8581 proper calculation;
wenzelm
parents: 6763
diff changeset
    71
qed;
6763
df12aef00932 Some bits of group theory. Demonstrate calculational proofs.
wenzelm
parents:
diff changeset
    72
6784
687ddcad8581 proper calculation;
wenzelm
parents: 6763
diff changeset
    73
text {*
7761
7fab9592384f improved presentation;
wenzelm
parents: 7748
diff changeset
    74
 \bigskip There are only two Isar language elements for calculational
7fab9592384f improved presentation;
wenzelm
parents: 7748
diff changeset
    75
 proofs: \isakeyword{also} for initial or intermediate calculational
7fab9592384f improved presentation;
wenzelm
parents: 7748
diff changeset
    76
 steps, and \isakeyword{finally} for building the result of a
7fab9592384f improved presentation;
wenzelm
parents: 7748
diff changeset
    77
 calculation.  These constructs are not hardwired into Isabelle/Isar,
7fab9592384f improved presentation;
wenzelm
parents: 7748
diff changeset
    78
 but defined on top of the basic Isar/VM interpreter.  Expanding the
7748
5b9c45b21782 improved presentation;
wenzelm
parents: 7740
diff changeset
    79
 \isakeyword{also} and \isakeyword{finally} derived language elements,
5b9c45b21782 improved presentation;
wenzelm
parents: 7740
diff changeset
    80
 calculations may be simulated as demonstrated below.
6784
687ddcad8581 proper calculation;
wenzelm
parents: 6763
diff changeset
    81
7748
5b9c45b21782 improved presentation;
wenzelm
parents: 7740
diff changeset
    82
 Note that ``$\dots$'' is just a special term binding that happens to
5b9c45b21782 improved presentation;
wenzelm
parents: 7740
diff changeset
    83
 be bound automatically to the argument of the last fact established
5b9c45b21782 improved presentation;
wenzelm
parents: 7740
diff changeset
    84
 by assume or any local goal.  In contrast to $\var{thesis}$, the
5b9c45b21782 improved presentation;
wenzelm
parents: 7740
diff changeset
    85
 ``$\dots$'' variable is bound \emph{after} the proof is finished.
6784
687ddcad8581 proper calculation;
wenzelm
parents: 6763
diff changeset
    86
*};
687ddcad8581 proper calculation;
wenzelm
parents: 6763
diff changeset
    87
6793
88aba7f486cb groups as monoids;
wenzelm
parents: 6784
diff changeset
    88
theorem "x * one = (x::'a::group)";
7133
64c9f2364dae renamed 'same' to '-';
wenzelm
parents: 7005
diff changeset
    89
proof -;
6784
687ddcad8581 proper calculation;
wenzelm
parents: 6763
diff changeset
    90
  have "x * one = x * (inv x * x)";
6908
1bf0590f4790 simp only;
wenzelm
parents: 6892
diff changeset
    91
    by (simp only: group_left_inverse);
6763
df12aef00932 Some bits of group theory. Demonstrate calculational proofs.
wenzelm
parents:
diff changeset
    92
7433
27887c52b9c8 "this";
wenzelm
parents: 7356
diff changeset
    93
  note calculation = this
6793
88aba7f486cb groups as monoids;
wenzelm
parents: 6784
diff changeset
    94
    -- {* first calculational step: init calculation register *};
6763
df12aef00932 Some bits of group theory. Demonstrate calculational proofs.
wenzelm
parents:
diff changeset
    95
6784
687ddcad8581 proper calculation;
wenzelm
parents: 6763
diff changeset
    96
  have "... = x * inv x * x";
6908
1bf0590f4790 simp only;
wenzelm
parents: 6892
diff changeset
    97
    by (simp only: group_assoc);
6763
df12aef00932 Some bits of group theory. Demonstrate calculational proofs.
wenzelm
parents:
diff changeset
    98
7433
27887c52b9c8 "this";
wenzelm
parents: 7356
diff changeset
    99
  note calculation = trans [OF calculation this]
6784
687ddcad8581 proper calculation;
wenzelm
parents: 6763
diff changeset
   100
    -- {* general calculational step: compose with transitivity rule *};
6763
df12aef00932 Some bits of group theory. Demonstrate calculational proofs.
wenzelm
parents:
diff changeset
   101
df12aef00932 Some bits of group theory. Demonstrate calculational proofs.
wenzelm
parents:
diff changeset
   102
  have "... = one * x";
6908
1bf0590f4790 simp only;
wenzelm
parents: 6892
diff changeset
   103
    by (simp only: group_right_inverse);
6763
df12aef00932 Some bits of group theory. Demonstrate calculational proofs.
wenzelm
parents:
diff changeset
   104
7433
27887c52b9c8 "this";
wenzelm
parents: 7356
diff changeset
   105
  note calculation = trans [OF calculation this]
6784
687ddcad8581 proper calculation;
wenzelm
parents: 6763
diff changeset
   106
    -- {* general calculational step: compose with transitivity rule *};
6763
df12aef00932 Some bits of group theory. Demonstrate calculational proofs.
wenzelm
parents:
diff changeset
   107
df12aef00932 Some bits of group theory. Demonstrate calculational proofs.
wenzelm
parents:
diff changeset
   108
  have "... = x";
6908
1bf0590f4790 simp only;
wenzelm
parents: 6892
diff changeset
   109
    by (simp only: group_left_unit);
6763
df12aef00932 Some bits of group theory. Demonstrate calculational proofs.
wenzelm
parents:
diff changeset
   110
7433
27887c52b9c8 "this";
wenzelm
parents: 7356
diff changeset
   111
  note calculation = trans [OF calculation this]
6784
687ddcad8581 proper calculation;
wenzelm
parents: 6763
diff changeset
   112
    -- {* final calculational step: compose with transitivity rule ... *};
687ddcad8581 proper calculation;
wenzelm
parents: 6763
diff changeset
   113
  from calculation
687ddcad8581 proper calculation;
wenzelm
parents: 6763
diff changeset
   114
    -- {* ... and pick up final result *};
6763
df12aef00932 Some bits of group theory. Demonstrate calculational proofs.
wenzelm
parents:
diff changeset
   115
7480
0a0e0dbe1269 replaced ?? by ?;
wenzelm
parents: 7433
diff changeset
   116
  show ?thesis; .;
6763
df12aef00932 Some bits of group theory. Demonstrate calculational proofs.
wenzelm
parents:
diff changeset
   117
qed;
df12aef00932 Some bits of group theory. Demonstrate calculational proofs.
wenzelm
parents:
diff changeset
   118
df12aef00932 Some bits of group theory. Demonstrate calculational proofs.
wenzelm
parents:
diff changeset
   119
7761
7fab9592384f improved presentation;
wenzelm
parents: 7748
diff changeset
   120
subsection {* Groups as monoids *};
6793
88aba7f486cb groups as monoids;
wenzelm
parents: 6784
diff changeset
   121
88aba7f486cb groups as monoids;
wenzelm
parents: 6784
diff changeset
   122
text {*
88aba7f486cb groups as monoids;
wenzelm
parents: 6784
diff changeset
   123
  Monoids are usually defined like this.
88aba7f486cb groups as monoids;
wenzelm
parents: 6784
diff changeset
   124
*};
88aba7f486cb groups as monoids;
wenzelm
parents: 6784
diff changeset
   125
88aba7f486cb groups as monoids;
wenzelm
parents: 6784
diff changeset
   126
axclass monoid < times
88aba7f486cb groups as monoids;
wenzelm
parents: 6784
diff changeset
   127
  monoid_assoc:       "(x * y) * z = x * (y * z)"
88aba7f486cb groups as monoids;
wenzelm
parents: 6784
diff changeset
   128
  monoid_left_unit:   "one * x = x"
88aba7f486cb groups as monoids;
wenzelm
parents: 6784
diff changeset
   129
  monoid_right_unit:  "x * one = x";
88aba7f486cb groups as monoids;
wenzelm
parents: 6784
diff changeset
   130
88aba7f486cb groups as monoids;
wenzelm
parents: 6784
diff changeset
   131
text {*
7748
5b9c45b21782 improved presentation;
wenzelm
parents: 7740
diff changeset
   132
 Groups are \emph{not} yet monoids directly from the definition.  For
5b9c45b21782 improved presentation;
wenzelm
parents: 7740
diff changeset
   133
 monoids, \name{right-unit} had to be included as an axiom, but for
5b9c45b21782 improved presentation;
wenzelm
parents: 7740
diff changeset
   134
 groups both \name{right-unit} and \name{right-inverse} are
5b9c45b21782 improved presentation;
wenzelm
parents: 7740
diff changeset
   135
 derivable from the other axioms.  With \name{group-right-unit}
5b9c45b21782 improved presentation;
wenzelm
parents: 7740
diff changeset
   136
 derived as a theorem of group theory (see above), we may still
5b9c45b21782 improved presentation;
wenzelm
parents: 7740
diff changeset
   137
 instantiate $\idt{group} \subset \idt{monoid}$ properly as follows.
6793
88aba7f486cb groups as monoids;
wenzelm
parents: 6784
diff changeset
   138
*};
88aba7f486cb groups as monoids;
wenzelm
parents: 6784
diff changeset
   139
88aba7f486cb groups as monoids;
wenzelm
parents: 6784
diff changeset
   140
instance group < monoid;
7356
1714c91b8729 expand_classes renamed to intro_classes;
wenzelm
parents: 7133
diff changeset
   141
  by (intro_classes,
6793
88aba7f486cb groups as monoids;
wenzelm
parents: 6784
diff changeset
   142
       rule group_assoc,
88aba7f486cb groups as monoids;
wenzelm
parents: 6784
diff changeset
   143
       rule group_left_unit,
88aba7f486cb groups as monoids;
wenzelm
parents: 6784
diff changeset
   144
       rule group_right_unit);
88aba7f486cb groups as monoids;
wenzelm
parents: 6784
diff changeset
   145
6763
df12aef00932 Some bits of group theory. Demonstrate calculational proofs.
wenzelm
parents:
diff changeset
   146
end;