src/HOL/AxClasses/Group/Group.thy
author clasohm
Wed, 20 Mar 1996 13:21:12 +0100
changeset 1589 fd6a571cb2b0
parent 1247 18b1441fb603
child 2907 0e272e4c7cb2
permissions -rw-r--r--
added warning and automatic deactivation of HTML generation if we cannot write .theory_list.txt; fixed bug which occured when index_path's value is "/"
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1247
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
     1
(*  Title:      Group.thy
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
     2
    ID:         $Id$
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
     3
    Author:     Markus Wenzel, TU Muenchen
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
     4
*)
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
     5
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
     6
Group = Sigs + Set +
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
     7
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
     8
(* semigroups *)
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
     9
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
    10
axclass
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
    11
  semigroup < times
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
    12
  assoc         "(x * y) * z = x * (y * z)"
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
    13
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
    14
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
    15
(* groups *)
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
    16
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
    17
axclass
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
    18
  group < one, inv, semigroup
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
    19
  left_unit     "1 * x = x"
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
    20
  left_inv      "inv x * x = 1"
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
    21
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
    22
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
    23
(* abelian groups *)
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
    24
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
    25
axclass
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
    26
  agroup < group
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
    27
  commut        "x * y = y * x"
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
    28
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
    29
end