File: fun_objects.v

package info (click to toggle)
coq 8.0pl2-2
  • links: PTS
  • area: main
  • in suites: sarge
  • size: 14,228 kB
  • ctags: 17,685
  • sloc: ml: 97,070; makefile: 1,255; sh: 738; lisp: 456; awk: 15
file content (32 lines) | stat: -rw-r--r-- 387 bytes parent folder | download | duplicates (2)
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
Implicit Arguments On.

Module Type SIG.
  Parameter id:(A:Set)A->A.
End SIG.
 
Module M[X:SIG].
  Definition idid := (X.id X.id).
  Definition id := (idid X.id).
End M.

Module N:=M.

Module Nat.
  Definition T := nat.
  Definition x := O.
  Definition id := [A:Set][x:A]x.
End Nat.

Module Z:=(N Nat).

Check (Z.idid O).

Module P[Y:SIG] := N.

Module Y:=P Nat Z.

Check (Y.id O).