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 33 34 35 36 37 38 39
|
Module Example1.
CoInductive wrap : Type :=
| item : unit -> wrap.
Definition extract (t : wrap) : unit :=
match t with
| item x => x
end.
CoFixpoint close u : unit -> wrap :=
match u with
| tt => item
end.
Definition table : wrap := close tt tt.
Eval vm_compute in (extract table).
Eval vm_compute in (extract table).
End Example1.
Module Example2.
Set Primitive Projections.
CoInductive wrap : Type :=
item { extract : unit }.
CoFixpoint close u : unit -> wrap :=
match u with
| tt => item
end.
Definition table : wrap := close tt tt.
Eval vm_compute in (extract table).
Eval vm_compute in (extract table).
End Example2.
|