| author | wenzelm | 
| Fri, 02 May 2014 19:28:32 +0200 | |
| changeset 56828 | 08041569357e | 
| parent 41460 | ea56b98aee83 | 
| child 58889 | 5b7a9633cfa8 | 
| permissions | -rw-r--r-- | 
| 5199 | 1  | 
(* Title: HOL/ex/MonoidGroup.thy  | 
2  | 
Author: Markus Wenzel  | 
|
3  | 
*)  | 
|
4  | 
||
| 17388 | 5  | 
header {* Monoids and Groups as predicates over record schemes *}
 | 
| 10357 | 6  | 
|
| 16417 | 7  | 
theory MonoidGroup imports Main begin  | 
| 5199 | 8  | 
|
9  | 
record 'a monoid_sig =  | 
|
| 11019 | 10  | 
times :: "'a => 'a => 'a"  | 
| 5199 | 11  | 
one :: 'a  | 
12  | 
||
| 9297 | 13  | 
record 'a group_sig = "'a monoid_sig" +  | 
| 5199 | 14  | 
inv :: "'a => 'a"  | 
15  | 
||
| 19736 | 16  | 
definition  | 
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
19736 
diff
changeset
 | 
17  | 
monoid :: "(| times :: 'a => 'a => 'a, one :: 'a, ... :: 'b |) => bool" where  | 
| 19736 | 18  | 
"monoid M = (\<forall>x y z.  | 
| 11019 | 19  | 
times M (times M x y) z = times M x (times M y z) \<and>  | 
| 19736 | 20  | 
times M (one M) x = x \<and> times M x (one M) = x)"  | 
| 5199 | 21  | 
|
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
19736 
diff
changeset
 | 
22  | 
definition  | 
| 
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
19736 
diff
changeset
 | 
23  | 
group :: "(| times :: 'a => 'a => 'a, one :: 'a, inv :: 'a => 'a, ... :: 'b |) => bool" where  | 
| 19736 | 24  | 
"group G = (monoid G \<and> (\<forall>x. times G (inv G x) x = one G))"  | 
| 5199 | 25  | 
|
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
19736 
diff
changeset
 | 
26  | 
definition  | 
| 11939 | 27  | 
reverse :: "(| times :: 'a => 'a => 'a, one :: 'a, ... :: 'b |) =>  | 
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
19736 
diff
changeset
 | 
28  | 
(| times :: 'a => 'a => 'a, one :: 'a, ... :: 'b |)" where  | 
| 19736 | 29  | 
"reverse M = M (| times := \<lambda>x y. times M y x |)"  | 
| 5199 | 30  | 
|
31  | 
end  |