2017-08-02 14:14:14 +02:00
|
|
|
|
Require Import HoTT HitTactics.
|
2017-08-03 15:07:53 +02:00
|
|
|
|
Require Import lattice representations.definition fsets.operations extensionality Sub fsets.properties.
|
2017-08-02 14:14:14 +02:00
|
|
|
|
|
|
|
|
|
Section k_finite.
|
|
|
|
|
|
2017-08-03 15:10:01 +02:00
|
|
|
|
Context (A : Type).
|
2017-08-02 14:14:14 +02:00
|
|
|
|
Context `{Univalence}.
|
|
|
|
|
|
2017-08-08 17:44:27 +02:00
|
|
|
|
Definition map (X : FSet A) : Sub A := fun a => a ∈ X.
|
2017-08-02 14:14:14 +02:00
|
|
|
|
|
2017-08-03 23:21:43 +02:00
|
|
|
|
Global Instance map_injective : IsEmbedding map.
|
2017-08-02 14:14:14 +02:00
|
|
|
|
Proof.
|
|
|
|
|
apply isembedding_isinj_hset. (* We use the fact that both [FSet A] and [Sub A] are hSets *)
|
|
|
|
|
intros X Y HXY.
|
|
|
|
|
apply fset_ext.
|
|
|
|
|
apply apD10. exact HXY.
|
|
|
|
|
Defined.
|
|
|
|
|
|
|
|
|
|
Definition Kf_sub_intern (B : Sub A) := exists (X : FSet A), B = map X.
|
|
|
|
|
|
|
|
|
|
Instance Kf_sub_hprop B : IsHProp (Kf_sub_intern B).
|
|
|
|
|
Proof.
|
|
|
|
|
apply hprop_allpath.
|
|
|
|
|
intros [X PX] [Y PY].
|
|
|
|
|
assert (X = Y) as HXY.
|
|
|
|
|
{ apply fset_ext. apply apD10.
|
|
|
|
|
transitivity B; [ symmetry | ]; assumption. }
|
|
|
|
|
apply path_sigma with HXY. simpl.
|
|
|
|
|
apply set_path2.
|
|
|
|
|
Defined.
|
|
|
|
|
|
|
|
|
|
Definition Kf_sub (B : Sub A) : hProp := BuildhProp (Kf_sub_intern B).
|
|
|
|
|
|
|
|
|
|
Definition Kf : hProp := Kf_sub (fun x => True).
|
|
|
|
|
|
2017-08-03 15:10:01 +02:00
|
|
|
|
Instance: IsHProp {X : FSet A & forall a : A, map X a}.
|
2017-08-02 14:14:14 +02:00
|
|
|
|
Proof.
|
2017-08-03 15:10:01 +02:00
|
|
|
|
apply hprop_allpath.
|
|
|
|
|
intros [X PX] [Y PY].
|
|
|
|
|
assert (X = Y) as HXY.
|
|
|
|
|
{ apply fset_ext. intros a.
|
|
|
|
|
unfold map in *.
|
|
|
|
|
apply path_hprop.
|
|
|
|
|
apply equiv_iff_hprop; intros.
|
|
|
|
|
+ apply PY.
|
|
|
|
|
+ apply PX. }
|
|
|
|
|
apply path_sigma with HXY. simpl.
|
|
|
|
|
apply path_forall. intro.
|
|
|
|
|
apply path_ishprop.
|
|
|
|
|
Defined.
|
|
|
|
|
|
|
|
|
|
Lemma Kf_unfold : Kf <~> (exists (X : FSet A), forall (a : A), map X a).
|
|
|
|
|
Proof.
|
|
|
|
|
apply equiv_equiv_iff_hprop. apply _. apply _.
|
2017-08-02 14:14:14 +02:00
|
|
|
|
split.
|
|
|
|
|
- intros [X PX]. exists X. intro a.
|
|
|
|
|
rewrite <- PX. done.
|
|
|
|
|
- intros [X PX]. exists X. apply path_forall; intro a.
|
|
|
|
|
apply path_hprop.
|
|
|
|
|
symmetry. apply if_hprop_then_equiv_Unit; [ apply _ | ].
|
|
|
|
|
apply PX.
|
|
|
|
|
Defined.
|
|
|
|
|
|
|
|
|
|
End k_finite.
|
2017-08-03 15:07:53 +02:00
|
|
|
|
|
2017-08-03 15:10:01 +02:00
|
|
|
|
Arguments map {_} {_} _.
|
|
|
|
|
|
2017-08-03 15:07:53 +02:00
|
|
|
|
Section structure_k_finite.
|
2017-08-03 15:10:01 +02:00
|
|
|
|
Context (A : Type).
|
2017-08-03 15:07:53 +02:00
|
|
|
|
Context `{Univalence}.
|
|
|
|
|
|
2017-08-08 17:44:27 +02:00
|
|
|
|
Lemma map_union : forall X Y : FSet A, map (X ∪ Y) = max_fun (map X) (map Y).
|
2017-08-03 15:07:53 +02:00
|
|
|
|
Proof.
|
|
|
|
|
intros.
|
|
|
|
|
unfold map, max_fun.
|
|
|
|
|
reflexivity.
|
|
|
|
|
Defined.
|
2017-08-03 15:10:01 +02:00
|
|
|
|
|
2017-08-08 17:44:27 +02:00
|
|
|
|
Lemma k_finite_union : closedUnion (Kf_sub A).
|
2017-08-03 15:07:53 +02:00
|
|
|
|
Proof.
|
2017-08-08 17:44:27 +02:00
|
|
|
|
unfold closedUnion, Kf_sub, Kf_sub_intern.
|
2017-08-03 15:07:53 +02:00
|
|
|
|
intros.
|
|
|
|
|
destruct X0 as [SX XP].
|
|
|
|
|
destruct X1 as [SY YP].
|
2017-08-08 17:44:27 +02:00
|
|
|
|
exists (SX ∪ SY).
|
2017-08-03 15:07:53 +02:00
|
|
|
|
rewrite map_union.
|
|
|
|
|
rewrite XP, YP.
|
|
|
|
|
reflexivity.
|
|
|
|
|
Defined.
|
|
|
|
|
|
2017-08-08 17:44:27 +02:00
|
|
|
|
Lemma k_finite_empty : closedEmpty (Kf_sub A).
|
2017-08-03 15:07:53 +02:00
|
|
|
|
Proof.
|
2017-08-08 17:44:27 +02:00
|
|
|
|
exists ∅.
|
2017-08-03 15:07:53 +02:00
|
|
|
|
reflexivity.
|
|
|
|
|
Defined.
|
|
|
|
|
|
2017-08-08 17:44:27 +02:00
|
|
|
|
Lemma k_finite_singleton : closedSingleton (Kf_sub A).
|
2017-08-03 15:07:53 +02:00
|
|
|
|
Proof.
|
|
|
|
|
intro.
|
2017-08-08 17:44:27 +02:00
|
|
|
|
exists {|a|}.
|
2017-08-03 15:07:53 +02:00
|
|
|
|
cbn.
|
|
|
|
|
apply path_forall.
|
2017-08-09 18:05:58 +02:00
|
|
|
|
intro z.
|
2017-08-03 15:07:53 +02:00
|
|
|
|
reflexivity.
|
|
|
|
|
Defined.
|
|
|
|
|
|
2017-08-03 15:10:01 +02:00
|
|
|
|
Lemma k_finite_hasDecidableEmpty : hasDecidableEmpty (Kf_sub A).
|
2017-08-03 15:07:53 +02:00
|
|
|
|
Proof.
|
2017-08-08 17:44:27 +02:00
|
|
|
|
unfold hasDecidableEmpty, closedEmpty, Kf_sub, Kf_sub_intern, map.
|
2017-08-03 15:07:53 +02:00
|
|
|
|
intros.
|
|
|
|
|
destruct X0 as [SX EX].
|
|
|
|
|
rewrite EX.
|
|
|
|
|
simple refine (Trunc_ind _ _ (merely_choice SX)).
|
|
|
|
|
intros [SXE | H1].
|
|
|
|
|
- rewrite SXE.
|
|
|
|
|
apply (tr (inl idpath)).
|
|
|
|
|
- apply (tr (inr H1)).
|
|
|
|
|
Defined.
|
|
|
|
|
End structure_k_finite.
|