2017-08-01 15:12:59 +02:00
|
|
|
|
(** Extensionality of the FSets *)
|
|
|
|
|
Require Import HoTT HitTactics.
|
2017-09-07 15:19:48 +02:00
|
|
|
|
Require Import kuratowski.kuratowski_sets.
|
2017-08-01 15:12:59 +02:00
|
|
|
|
|
|
|
|
|
Section ext.
|
2017-08-02 11:40:03 +02:00
|
|
|
|
Context {A : Type}.
|
|
|
|
|
Context `{Univalence}.
|
2017-08-01 15:12:59 +02:00
|
|
|
|
|
2017-08-14 21:38:50 +02:00
|
|
|
|
Lemma equiv_subset1_l (X Y : FSet A) (H1 : Y ∪ X = X) (a : A) (Ya : a ∈ Y) : a ∈ X.
|
|
|
|
|
Proof.
|
|
|
|
|
apply (transport (fun Z => a ∈ Z) H1 (tr(inl Ya))).
|
|
|
|
|
Defined.
|
|
|
|
|
|
|
|
|
|
Lemma equiv_subset1_r X : forall (Y : FSet A), (forall a, a ∈ Y -> a ∈ X) -> Y ∪ X = X.
|
|
|
|
|
Proof.
|
|
|
|
|
hinduction ; try (intros ; apply path_ishprop).
|
|
|
|
|
- intros.
|
|
|
|
|
apply nl.
|
|
|
|
|
- intros b sub.
|
|
|
|
|
specialize (sub b (tr idpath)).
|
|
|
|
|
revert sub.
|
|
|
|
|
hinduction X ; try (intros ; apply path_ishprop).
|
|
|
|
|
* contradiction.
|
|
|
|
|
* intros.
|
|
|
|
|
strip_truncations.
|
|
|
|
|
rewrite sub.
|
|
|
|
|
apply union_idem.
|
|
|
|
|
* intros X Y subX subY mem.
|
|
|
|
|
strip_truncations.
|
|
|
|
|
destruct mem as [t | t].
|
|
|
|
|
** rewrite assoc, (subX t).
|
|
|
|
|
reflexivity.
|
|
|
|
|
** rewrite (comm X), assoc, (subY t).
|
|
|
|
|
reflexivity.
|
|
|
|
|
- intros Y1 Y2 H1 H2 H3.
|
|
|
|
|
rewrite <- assoc.
|
|
|
|
|
rewrite (H2 (fun a HY => H3 a (tr(inr HY)))).
|
|
|
|
|
apply (H1 (fun a HY => H3 a (tr(inl HY)))).
|
|
|
|
|
Defined.
|
|
|
|
|
|
|
|
|
|
Lemma eq_subset1 X Y : (Y ∪ X = X) * (X ∪ Y = Y) <~> forall (a : A), a ∈ X = a ∈ Y.
|
|
|
|
|
Proof.
|
|
|
|
|
eapply equiv_iff_hprop_uncurried ; split.
|
|
|
|
|
- intros [H1 H2] a.
|
|
|
|
|
apply path_iff_hprop ; apply equiv_subset1_l ; assumption.
|
|
|
|
|
- intros H1.
|
|
|
|
|
split ; apply equiv_subset1_r ; intros.
|
|
|
|
|
* rewrite H1 ; assumption.
|
|
|
|
|
* rewrite <- H1 ; assumption.
|
|
|
|
|
Defined.
|
|
|
|
|
|
|
|
|
|
Lemma eq_subset2 (X Y : FSet A) : X = Y <~> (Y ∪ X = X) * (X ∪ Y = Y).
|
|
|
|
|
Proof.
|
|
|
|
|
eapply equiv_iff_hprop_uncurried ; split.
|
|
|
|
|
- intro Heq.
|
|
|
|
|
split.
|
|
|
|
|
* apply (ap (fun Z => Z ∪ X) Heq^ @ union_idem X).
|
|
|
|
|
* apply (ap (fun Z => Z ∪ Y) Heq @ union_idem Y).
|
|
|
|
|
- intros [H1 H2].
|
|
|
|
|
apply (H1^ @ comm Y X @ H2).
|
|
|
|
|
Defined.
|
|
|
|
|
|
|
|
|
|
Theorem fset_ext (X Y : FSet A) :
|
|
|
|
|
X = Y <~> forall (a : A), a ∈ X = a ∈ Y.
|
|
|
|
|
Proof.
|
|
|
|
|
apply (equiv_compose' (eq_subset1 X Y) (eq_subset2 X Y)).
|
|
|
|
|
Defined.
|
|
|
|
|
|
2017-08-09 18:05:58 +02:00
|
|
|
|
Lemma subset_union (X Y : FSet A) :
|
2017-08-08 13:35:28 +02:00
|
|
|
|
X ⊆ Y -> X ∪ Y = Y.
|
|
|
|
|
Proof.
|
2017-08-14 21:38:50 +02:00
|
|
|
|
hinduction X ; try (intros ; apply path_ishprop).
|
|
|
|
|
- intros.
|
|
|
|
|
apply nl.
|
2017-08-08 13:35:28 +02:00
|
|
|
|
- intros a.
|
2017-08-14 21:38:50 +02:00
|
|
|
|
hinduction Y ; try (intros ; apply path_ishprop).
|
2017-08-08 13:35:28 +02:00
|
|
|
|
+ intro.
|
|
|
|
|
contradiction.
|
2017-08-14 21:38:50 +02:00
|
|
|
|
+ intros b p.
|
|
|
|
|
strip_truncations.
|
|
|
|
|
rewrite p.
|
|
|
|
|
apply idem.
|
|
|
|
|
+ intros X1 X2 IH1 IH2 t.
|
|
|
|
|
strip_truncations.
|
|
|
|
|
destruct t as [t | t].
|
|
|
|
|
++ rewrite assoc, (IH1 t).
|
2017-08-08 13:35:28 +02:00
|
|
|
|
reflexivity.
|
2017-08-14 21:38:50 +02:00
|
|
|
|
++ rewrite comm, <- assoc, (comm X2), (IH2 t).
|
2017-08-08 13:35:28 +02:00
|
|
|
|
reflexivity.
|
|
|
|
|
- intros X1 X2 IH1 IH2 [G1 G2].
|
|
|
|
|
rewrite <- assoc.
|
|
|
|
|
rewrite (IH2 G2).
|
|
|
|
|
apply (IH1 G1).
|
|
|
|
|
Defined.
|
|
|
|
|
|
|
|
|
|
Lemma subset_union_l (X : FSet A) :
|
|
|
|
|
forall Y, X ⊆ X ∪ Y.
|
|
|
|
|
Proof.
|
2017-08-14 21:38:50 +02:00
|
|
|
|
hinduction X ; try (intros ; apply path_ishprop).
|
2017-08-08 13:35:28 +02:00
|
|
|
|
- apply (fun _ => tt).
|
2017-08-14 21:38:50 +02:00
|
|
|
|
- intros.
|
2017-08-08 13:35:28 +02:00
|
|
|
|
apply (tr(inl(tr idpath))).
|
|
|
|
|
- intros X1 X2 HX1 HX2 Y.
|
2017-08-08 17:00:30 +02:00
|
|
|
|
split ; unfold subset in *.
|
2017-08-14 21:38:50 +02:00
|
|
|
|
* rewrite <- assoc.
|
|
|
|
|
apply HX1.
|
|
|
|
|
* rewrite (comm X1 X2), <- assoc.
|
|
|
|
|
apply HX2.
|
2017-08-08 13:35:28 +02:00
|
|
|
|
Defined.
|
|
|
|
|
|
2017-08-02 11:40:03 +02:00
|
|
|
|
Lemma subset_union_equiv
|
2017-08-07 16:49:46 +02:00
|
|
|
|
: forall X Y : FSet A, X ⊆ Y <~> X ∪ Y = Y.
|
2017-08-02 11:40:03 +02:00
|
|
|
|
Proof.
|
|
|
|
|
intros X Y.
|
|
|
|
|
eapply equiv_iff_hprop_uncurried.
|
|
|
|
|
split.
|
|
|
|
|
- apply subset_union.
|
|
|
|
|
- intro HXY.
|
|
|
|
|
rewrite <- HXY.
|
|
|
|
|
apply subset_union_l.
|
|
|
|
|
Defined.
|
2017-08-14 21:38:50 +02:00
|
|
|
|
|
2017-08-02 11:40:03 +02:00
|
|
|
|
Lemma subset_isIn (X Y : FSet A) :
|
2017-08-14 21:38:50 +02:00
|
|
|
|
X ⊆ Y <~> forall (a : A), a ∈ X -> a ∈ Y.
|
2017-08-02 11:40:03 +02:00
|
|
|
|
Proof.
|
2017-08-14 21:38:50 +02:00
|
|
|
|
etransitivity.
|
|
|
|
|
- apply subset_union_equiv.
|
|
|
|
|
- eapply equiv_iff_hprop_uncurried ; split.
|
|
|
|
|
* apply equiv_subset1_l.
|
|
|
|
|
* apply equiv_subset1_r.
|
2017-08-02 11:40:03 +02:00
|
|
|
|
Defined.
|
2017-08-01 15:12:59 +02:00
|
|
|
|
|
2017-08-02 11:40:03 +02:00
|
|
|
|
Lemma eq_subset (X Y : FSet A) :
|
2017-08-07 16:49:46 +02:00
|
|
|
|
X = Y <~> (Y ⊆ X * X ⊆ Y).
|
2017-08-02 11:40:03 +02:00
|
|
|
|
Proof.
|
2017-08-14 21:38:50 +02:00
|
|
|
|
etransitivity ((Y ∪ X = X) * (X ∪ Y = Y)).
|
|
|
|
|
- apply eq_subset2.
|
|
|
|
|
- symmetry.
|
|
|
|
|
eapply equiv_functor_prod' ; apply subset_union_equiv.
|
2017-08-02 11:40:03 +02:00
|
|
|
|
Defined.
|
2017-08-14 21:38:50 +02:00
|
|
|
|
End ext.
|