2017-09-07 15:19:48 +02:00
|
|
|
|
(** Operations on the [FSet A] for an arbitrary [A] *)
|
2017-09-21 14:12:51 +02:00
|
|
|
|
Require Import HoTT HitTactics prelude.
|
|
|
|
|
Require Import kuratowski_sets monad_interface extensionality
|
|
|
|
|
list_representation.isomorphism list_representation.list_representation.
|
2017-05-23 16:30:31 +02:00
|
|
|
|
|
2017-08-01 15:12:59 +02:00
|
|
|
|
Section operations.
|
2017-09-07 15:19:48 +02:00
|
|
|
|
(** Monad operations. *)
|
|
|
|
|
Definition fset_fmap {A B : Type} : (A -> B) -> FSet A -> FSet B.
|
|
|
|
|
Proof.
|
|
|
|
|
intro f.
|
|
|
|
|
hrecursion.
|
|
|
|
|
- exact ∅.
|
|
|
|
|
- apply (fun a => {|f a|}).
|
|
|
|
|
- apply (∪).
|
|
|
|
|
- apply assoc.
|
|
|
|
|
- apply comm.
|
|
|
|
|
- apply nl.
|
|
|
|
|
- apply nr.
|
|
|
|
|
- apply (idem o f).
|
|
|
|
|
Defined.
|
|
|
|
|
|
2017-11-01 17:47:41 +01:00
|
|
|
|
Global Instance return_fset : Return FSet := fun _ a => {|a|}.
|
2017-05-23 16:30:31 +02:00
|
|
|
|
|
2017-11-01 17:47:41 +01:00
|
|
|
|
Global Instance bind_fset : Bind FSet.
|
2017-08-07 16:49:46 +02:00
|
|
|
|
Proof.
|
2017-11-01 17:47:41 +01:00
|
|
|
|
intros A B m f.
|
|
|
|
|
hrecursion m.
|
2017-09-07 15:19:48 +02:00
|
|
|
|
- exact ∅.
|
2017-11-01 17:47:41 +01:00
|
|
|
|
- exact f.
|
2017-09-07 15:19:48 +02:00
|
|
|
|
- apply (∪).
|
|
|
|
|
- apply assoc.
|
|
|
|
|
- apply comm.
|
|
|
|
|
- apply nl.
|
|
|
|
|
- apply nr.
|
2017-11-01 17:47:41 +01:00
|
|
|
|
- intros; apply union_idem.
|
2017-08-07 16:49:46 +02:00
|
|
|
|
Defined.
|
2017-05-23 16:30:31 +02:00
|
|
|
|
|
2017-09-07 15:19:48 +02:00
|
|
|
|
(** Set-theoretical operations. *)
|
2017-08-08 17:00:30 +02:00
|
|
|
|
Global Instance fset_comprehension : forall A, hasComprehension (FSet A) A.
|
2017-08-07 16:49:46 +02:00
|
|
|
|
Proof.
|
2017-08-08 17:00:30 +02:00
|
|
|
|
intros A P.
|
2017-08-07 16:49:46 +02:00
|
|
|
|
hrecursion.
|
|
|
|
|
- apply ∅.
|
|
|
|
|
- intro a.
|
|
|
|
|
refine (if (P a) then {|a|} else ∅).
|
2017-08-08 15:29:50 +02:00
|
|
|
|
- apply (∪).
|
2017-08-07 16:49:46 +02:00
|
|
|
|
- apply assoc.
|
|
|
|
|
- apply comm.
|
|
|
|
|
- apply nl.
|
|
|
|
|
- apply nr.
|
2017-08-09 18:05:58 +02:00
|
|
|
|
- intros; simpl.
|
2017-10-11 17:11:07 +02:00
|
|
|
|
apply union_idem.
|
2017-08-07 16:49:46 +02:00
|
|
|
|
Defined.
|
2017-09-07 17:22:14 +02:00
|
|
|
|
|
2017-08-08 17:00:30 +02:00
|
|
|
|
Definition single_product {A B : Type} (a : A) : FSet B -> FSet (A * B).
|
2017-08-07 23:15:25 +02:00
|
|
|
|
Proof.
|
|
|
|
|
hrecursion.
|
|
|
|
|
- apply ∅.
|
|
|
|
|
- intro b.
|
|
|
|
|
apply {|(a, b)|}.
|
2017-08-08 15:29:50 +02:00
|
|
|
|
- apply (∪).
|
2017-09-07 15:19:48 +02:00
|
|
|
|
- apply assoc.
|
|
|
|
|
- apply comm.
|
|
|
|
|
- apply nl.
|
|
|
|
|
- apply nr.
|
|
|
|
|
- intros.
|
|
|
|
|
apply idem.
|
2017-08-07 23:15:25 +02:00
|
|
|
|
Defined.
|
2017-08-09 18:05:58 +02:00
|
|
|
|
|
2017-08-08 17:00:30 +02:00
|
|
|
|
Definition product {A B : Type} : FSet A -> FSet B -> FSet (A * B).
|
2017-08-07 22:13:42 +02:00
|
|
|
|
Proof.
|
|
|
|
|
intros X Y.
|
|
|
|
|
hrecursion X.
|
|
|
|
|
- apply ∅.
|
|
|
|
|
- intro a.
|
2017-08-07 23:15:25 +02:00
|
|
|
|
apply (single_product a Y).
|
2017-08-08 15:29:50 +02:00
|
|
|
|
- apply (∪).
|
2017-09-07 15:19:48 +02:00
|
|
|
|
- apply assoc.
|
|
|
|
|
- apply comm.
|
|
|
|
|
- apply nl.
|
|
|
|
|
- apply nr.
|
2017-08-08 13:35:28 +02:00
|
|
|
|
- intros ; apply union_idem.
|
2017-08-07 22:13:42 +02:00
|
|
|
|
Defined.
|
2017-08-09 18:05:58 +02:00
|
|
|
|
|
2017-09-07 17:22:14 +02:00
|
|
|
|
Definition disjoint_sum {A B : Type} (X : FSet A) (Y : FSet B) : FSet (A + B) :=
|
|
|
|
|
(fset_fmap inl X) ∪ (fset_fmap inr Y).
|
|
|
|
|
|
2017-08-09 17:03:51 +02:00
|
|
|
|
Local Ltac remove_transport
|
|
|
|
|
:= intros ; apply path_forall ; intro Z ; rewrite transport_arrow
|
|
|
|
|
; rewrite transport_const ; simpl.
|
|
|
|
|
Local Ltac pointwise
|
|
|
|
|
:= repeat (f_ap) ; try (apply path_forall ; intro Z2) ;
|
|
|
|
|
rewrite transport_arrow, transport_const, transport_sigma
|
|
|
|
|
; f_ap ; hott_simpl ; simple refine (path_sigma _ _ _ _ _)
|
|
|
|
|
; try (apply transport_const) ; try (apply path_ishprop).
|
|
|
|
|
|
2017-09-07 15:19:48 +02:00
|
|
|
|
Context `{Univalence}.
|
|
|
|
|
|
|
|
|
|
(** Separation axiom for finite sets. *)
|
|
|
|
|
Definition separation (A B : Type) : forall (X : FSet A) (f : {a | a ∈ X} -> B), FSet B.
|
2017-08-09 17:03:51 +02:00
|
|
|
|
Proof.
|
|
|
|
|
hinduction.
|
|
|
|
|
- intros f.
|
|
|
|
|
apply ∅.
|
|
|
|
|
- intros a f.
|
|
|
|
|
apply {|f (a; tr idpath)|}.
|
|
|
|
|
- intros X1 X2 HX1 HX2 f.
|
2017-11-01 17:47:41 +01:00
|
|
|
|
pose (fX1 := fun Z : {a : A | a ∈ X1} => f (pr1 Z;tr (inl (pr2 Z)))).
|
|
|
|
|
pose (fX2 := fun Z : {a : A | a ∈ X2} => f (pr1 Z;tr (inr (pr2 Z)))).
|
2017-08-09 17:03:51 +02:00
|
|
|
|
specialize (HX1 fX1).
|
|
|
|
|
specialize (HX2 fX2).
|
|
|
|
|
apply (HX1 ∪ HX2).
|
|
|
|
|
- remove_transport.
|
|
|
|
|
rewrite assoc.
|
|
|
|
|
pointwise.
|
|
|
|
|
- remove_transport.
|
|
|
|
|
rewrite comm.
|
|
|
|
|
pointwise.
|
|
|
|
|
- remove_transport.
|
|
|
|
|
rewrite nl.
|
|
|
|
|
pointwise.
|
|
|
|
|
- remove_transport.
|
|
|
|
|
rewrite nr.
|
|
|
|
|
pointwise.
|
|
|
|
|
- remove_transport.
|
|
|
|
|
rewrite <- (idem (Z (x; tr 1%path))).
|
|
|
|
|
pointwise.
|
2017-08-09 18:05:58 +02:00
|
|
|
|
Defined.
|
2017-09-07 17:22:14 +02:00
|
|
|
|
|
|
|
|
|
(** [FSet A] has decidable emptiness. *)
|
|
|
|
|
Definition isEmpty {A : Type} (X : FSet A) : Decidable (X = ∅).
|
|
|
|
|
Proof.
|
|
|
|
|
hinduction X ; try (intros ; apply path_ishprop).
|
|
|
|
|
- apply (inl idpath).
|
|
|
|
|
- intros.
|
|
|
|
|
refine (inr (fun p => _)).
|
|
|
|
|
refine (transport (fun Z : hProp => Z) (ap (fun z => a ∈ z) p) _).
|
|
|
|
|
apply (tr idpath).
|
|
|
|
|
- intros X Y H1 H2.
|
|
|
|
|
destruct H1 as [p1 | p1], H2 as [p2 | p2].
|
|
|
|
|
* left.
|
|
|
|
|
rewrite p1, p2.
|
|
|
|
|
apply nl.
|
|
|
|
|
* right.
|
|
|
|
|
rewrite p1, nl.
|
|
|
|
|
apply p2.
|
|
|
|
|
* right.
|
|
|
|
|
rewrite p2, nr.
|
|
|
|
|
apply p1.
|
|
|
|
|
* right.
|
|
|
|
|
intros p.
|
|
|
|
|
apply p1.
|
|
|
|
|
apply fset_ext.
|
|
|
|
|
intros.
|
|
|
|
|
apply path_iff_hprop ; try contradiction.
|
|
|
|
|
intro H1.
|
|
|
|
|
refine (transport (fun z => a ∈ z) p _).
|
|
|
|
|
apply (tr(inl H1)).
|
|
|
|
|
Defined.
|
2017-08-09 18:05:58 +02:00
|
|
|
|
End operations.
|
2017-09-07 15:19:48 +02:00
|
|
|
|
|
|
|
|
|
Section operations_decidable.
|
|
|
|
|
Context {A : Type}.
|
2017-09-21 14:12:51 +02:00
|
|
|
|
Context `{MerelyDecidablePaths A}.
|
2017-09-07 15:19:48 +02:00
|
|
|
|
Context `{Univalence}.
|
|
|
|
|
|
|
|
|
|
Global Instance isIn_decidable (a : A) : forall (X : FSet A),
|
|
|
|
|
Decidable (a ∈ X).
|
|
|
|
|
Proof.
|
|
|
|
|
hinduction ; try (intros ; apply path_ishprop).
|
|
|
|
|
- apply _.
|
2017-09-21 14:12:51 +02:00
|
|
|
|
- apply (m_dec_path _).
|
2017-09-07 15:19:48 +02:00
|
|
|
|
- apply _.
|
|
|
|
|
Defined.
|
|
|
|
|
|
2017-09-22 16:16:12 +02:00
|
|
|
|
Global Instance fset_member_bool : hasMembership_decidable (FSet A) A :=
|
|
|
|
|
fun a X =>
|
|
|
|
|
match (dec a ∈ X) with
|
|
|
|
|
| inl _ => true
|
|
|
|
|
| inr _ => false
|
|
|
|
|
end.
|
2017-09-07 15:19:48 +02:00
|
|
|
|
|
|
|
|
|
Global Instance subset_decidable : forall (X Y : FSet A), Decidable (X ⊆ Y).
|
|
|
|
|
Proof.
|
|
|
|
|
hinduction ; try (intros ; apply path_ishprop).
|
|
|
|
|
- apply _.
|
|
|
|
|
- apply _.
|
|
|
|
|
- unfold subset in *.
|
|
|
|
|
apply _.
|
|
|
|
|
Defined.
|
|
|
|
|
|
|
|
|
|
Global Instance fset_subset_bool : hasSubset_decidable (FSet A).
|
|
|
|
|
Proof.
|
|
|
|
|
intros X Y.
|
|
|
|
|
apply (if (dec (X ⊆ Y)) then true else false).
|
|
|
|
|
Defined.
|
|
|
|
|
|
|
|
|
|
Global Instance fset_intersection : hasIntersection (FSet A)
|
2017-11-01 17:47:41 +01:00
|
|
|
|
:= fun X Y => {|X | fun a => a ∈_d Y |}.
|
2017-09-21 18:09:40 +02:00
|
|
|
|
|
2017-11-01 17:47:41 +01:00
|
|
|
|
Definition difference := fun X Y => {|X | fun a => negb a ∈_d Y|}.
|
2017-09-21 14:12:51 +02:00
|
|
|
|
End operations_decidable.
|
|
|
|
|
|
|
|
|
|
Section FSet_cons_rec.
|
|
|
|
|
Context `{A : Type}.
|
|
|
|
|
|
|
|
|
|
Variable (P : Type)
|
|
|
|
|
(Pset : IsHSet P)
|
|
|
|
|
(Pe : P)
|
|
|
|
|
(Pcons : A -> FSet A -> P -> P)
|
|
|
|
|
(Pdupl : forall X a p, Pcons a ({|a|} ∪ X) (Pcons a X p) = Pcons a X p)
|
|
|
|
|
(Pcomm : forall X a b p, Pcons a ({|b|} ∪ X) (Pcons b X p)
|
|
|
|
|
= Pcons b ({|a|} ∪ X) (Pcons a X p)).
|
|
|
|
|
|
2017-09-26 11:56:58 +02:00
|
|
|
|
Definition FSet_cons_rec (X : FSet A) : P :=
|
|
|
|
|
FSetC_prim_rec A P Pset Pe
|
|
|
|
|
(fun a Y p => Pcons a (FSetC_to_FSet Y) p)
|
|
|
|
|
(fun _ _ => Pdupl _ _)
|
|
|
|
|
(fun _ _ _ => Pcomm _ _ _)
|
|
|
|
|
(FSet_to_FSetC X).
|
2017-09-21 14:12:51 +02:00
|
|
|
|
|
|
|
|
|
Definition FSet_cons_beta_empty : FSet_cons_rec ∅ = Pe := idpath.
|
|
|
|
|
|
|
|
|
|
Definition FSet_cons_beta_cons (a : A) (X : FSet A)
|
|
|
|
|
: FSet_cons_rec ({|a|} ∪ X) = Pcons a X (FSet_cons_rec X)
|
|
|
|
|
:= ap (fun z => Pcons a z _) (repr_iso_id_l _).
|
2017-09-24 23:35:45 +02:00
|
|
|
|
End FSet_cons_rec.
|