2017-09-07 15:19:48 +02:00
|
|
|
(** Operations on [FSetC A] *)
|
2017-08-02 11:40:03 +02:00
|
|
|
Require Import HoTT HitTactics.
|
2017-09-07 15:19:48 +02:00
|
|
|
Require Import list_representation.
|
2017-08-02 11:40:03 +02:00
|
|
|
|
|
|
|
Section operations.
|
2017-08-08 17:00:30 +02:00
|
|
|
Global Instance fsetc_union : forall A, hasUnion (FSetC A).
|
2017-08-08 15:29:50 +02:00
|
|
|
Proof.
|
|
|
|
intros A x y.
|
2017-08-02 11:40:03 +02:00
|
|
|
hinduction x.
|
|
|
|
- apply y.
|
|
|
|
- apply Cns.
|
|
|
|
- apply dupl.
|
2017-09-21 14:12:51 +02:00
|
|
|
- apply comm_s.
|
2017-08-02 11:40:03 +02:00
|
|
|
Defined.
|
|
|
|
|
2017-08-08 17:00:30 +02:00
|
|
|
Global Instance fsetc_singleton : forall A, hasSingleton (FSetC A) A := fun A a => a;;∅.
|
2017-08-02 11:40:03 +02:00
|
|
|
|
|
|
|
End operations.
|