src/HOL/ex/MonoidGroup.thy
author paulson
Thu, 01 Dec 2005 15:45:54 +0100
changeset 18315 e52f867ab851
parent 17388 495c799df31d
child 19736 d8d0f8f51d69
permissions -rw-r--r--
restoring the old status of subset_refl

(*  Title:      HOL/ex/MonoidGroup.thy
    ID:         $Id$
    Author:     Markus Wenzel
*)

header {* Monoids and Groups as predicates over record schemes *}

theory MonoidGroup imports Main begin

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 |) => 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 |) => bool"
  "group G == monoid G \<and> (\<forall>x. times G (inv G x) x = one G)"

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

end