1 2 3 4 5 6 7 8 9 10 11 12 13 14 15 16 17 18 19 20 21 22 23 24 25 26 27 28 29 30 31 32
|
Set Implicit Arguments.
Unset Strict Implicit.
Module M.
Definition id (A : Set) (x : A) := x.
Module Type SIG.
Parameter idid : forall A : Set, A -> A.
End SIG.
Module N.
Definition idid (A : Set) (x : A) := id x.
(* <Warning> : Grammar is replaced by Notation *)
Notation inc := (plus 1).
End N.
Definition zero := N.idid 0.
End M.
Definition zero := M.N.idid 0.
Definition jeden := M.N.inc 0.
Module Goly := M.N.
Definition Gole_zero := Goly.idid 0.
Definition Goly_jeden := Goly.inc 0.
Module Ubrany : M.SIG := M.N.
Definition Ubrane_zero := Ubrany.idid 0.
|