src/HOL/AxClasses/Group/GroupDefs.thy
author wenzelm
Mon, 29 Nov 1999 15:52:49 +0100
changeset 8039 a901bafe4578
parent 2907 0e272e4c7cb2
permissions -rw-r--r--
Goal: tuned pris;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1247
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
     1
(*  Title:      GroupDefs.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
GroupDefs = MonoidGroupInsts + Prod + Fun +
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
     7
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
     8
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
     9
(* bool *)
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
    10
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
    11
instance
2907
0e272e4c7cb2 inv -> inverse
nipkow
parents: 1247
diff changeset
    12
  bool :: {times, inverse, one}
1247
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
    13
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
    14
defs
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
    15
  times_bool_def        "x * y == (x ~= y)"
2907
0e272e4c7cb2 inv -> inverse
nipkow
parents: 1247
diff changeset
    16
  inverse_bool_def      "inverse x == (x::bool)"
1247
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
    17
  one_bool_def          "1 == False"
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
    18
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
    19
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
    20
(* cartesian products *)
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
    21
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
    22
instance
2907
0e272e4c7cb2 inv -> inverse
nipkow
parents: 1247
diff changeset
    23
  "*" :: (term, term) {times, inverse, one}
1247
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
    24
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
    25
defs
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
    26
  times_prod_def        "p * q == (fst p * fst q, snd p * snd q)"
2907
0e272e4c7cb2 inv -> inverse
nipkow
parents: 1247
diff changeset
    27
  inverse_prod_def      "inverse p == (inverse (fst p), inverse (snd p))"
1247
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
    28
  one_prod_def          "1 == (1, 1)"
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
    29
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
    30
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
    31
(* function spaces *)
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
    32
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
    33
instance
2907
0e272e4c7cb2 inv -> inverse
nipkow
parents: 1247
diff changeset
    34
  fun :: (term, term) {times, inverse, one}
1247
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
    35
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
    36
defs
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
    37
  times_fun_def         "f * g == (%x. f x * g x)"
2907
0e272e4c7cb2 inv -> inverse
nipkow
parents: 1247
diff changeset
    38
  inverse_fun_def       "inverse f == (%x. inverse (f x))"
1247
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
    39
  one_fun_def           "1 == (%x. 1)"
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
    40
18b1441fb603 Various axiomatic type class demos;
wenzelm
parents:
diff changeset
    41
end