src/HOL/ex/MonoidGroup.thy
author nipkow
Mon, 06 Aug 2001 13:43:24 +0200
changeset 11464 ddea204de5bc
parent 11019 e968e5bfe98d
child 11939 c5e69470f03b
permissions -rw-r--r--
turned translation for 1::nat into def. introduced 1' and replaced most occurrences of 1 by 1'.

(*  Title:      HOL/ex/MonoidGroup.thy
    ID:         $Id$
    Author:     Markus Wenzel
    License:    GPL (GNU GENERAL PUBLIC LICENSE)

Monoids and Groups as predicates over record schemes.
*)

header {* Monoids and Groups *}

theory MonoidGroup = Main:

record 'a monoid_sig =
  times :: "'a => 'a => 'a"
  one :: 'a

record 'a group_sig = "'a monoid_sig" +
  inv :: "'a => 'a"

constdefs
  monoid :: "(| times :: 'a => 'a => 'a, one :: 'a, ... :: 'b::more |) => bool"
  "monoid M == \<forall>x y z.
    times M (times M x y) z = times M x (times M y z) \<and>
    times M (one M) x = x \<and> times M x (one M) = x"

  group :: "(| times :: 'a => 'a => 'a, one :: 'a, inv :: 'a => 'a, ... :: 'b::more |) => bool"
  "group G == monoid G \<and> (\<forall>x. times G (inv G x) x = one G)"

  reverse :: "(| times :: 'a => 'a => 'a, one :: 'a, ... :: 'b::more |) =>
    (| times :: 'a => 'a => 'a, one :: 'a, ... :: 'b |)"
  "reverse M == M (| times := \<lambda>x y. times M y x |)"

end