1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| The HELM team. *)
8 (* ||A|| http://helm.cs.unibo.it *)
10 (* \ / This file is distributed under the terms of the *)
11 (* v GNU General Public License Version 2 *)
13 (**************************************************************************)
15 include "logic/equality.ma".
17 inductive unit (A : Type) (a : A) : Type := mk_unit : unit A a.
19 definition k : ∀A,a.unit A a → A := λA,a.λx:unit A a.a.
22 record monoid : Type := {
26 mP : ∀x:mC.x = mOp x m1
29 record group : Type := {
34 gP : ∀x:gC.gOp x (gInv x) = g1
38 record monoid_aux (x : Type) : Type := {
40 xOp : let xC := k ?? xC in xC -> xC -> xC;
41 x1 : let xC := k ?? xC in xC;
42 xP : let xC := k ?? xC in ∀x:xC. x = xOp x x1
46 record monoid_aux (xC : Type) : Type := {
48 xOp : (*let xC := k ?? xC in*) xC -> xC -> xC;
49 x1 : (*let xC := k ?? xC in*) xC;
50 xP : (*let xC := k ?? xC in*) ∀x:xC. x = xOp x x1
54 definition m_sub := λT:Type.λx:monoid_aux T.
55 mk_monoid (*k ?? (xC T x)*) T (xOp ? x) (x1 ? x) (xP ? x).
58 lemma test : ∀G:group.∀M:monoid_aux (gC G).gC G = mC (m_sub ? M).
62 record ring : Type := {
64 rM :> monoid_aux (gC rG);
65 rP : ∀x,y,z:rG. mOp rM x (gOp rG y z) = gOp rG (mOp rM x y) (mOp rM x z)
68 lemma prova : ∀R:ring.∀x:R. x = mOp R x (m1 R).
92 lemma em_1 : monoid_aux Z.
101 [ apply e_g; | apply em_1;| intros; unfold e_g; unfold em_1; simplify; autobatch;]
104 lemma test1 : ∀x:Z.x = x * pos O.
105 intros; apply (mP e_r x);