mirror of https://github.com/nmvdw/HITs-Examples
Finalize the definition of K-finite (sub)objects
This commit is contained in:
parent
fa076f5f41
commit
4141f9d456
|
@ -18,5 +18,6 @@ fsets/monad.v
|
||||||
FSets.v
|
FSets.v
|
||||||
implementations/lists.v
|
implementations/lists.v
|
||||||
variations/enumerated.v
|
variations/enumerated.v
|
||||||
|
variations/k_finite.v
|
||||||
#empty_set.v
|
#empty_set.v
|
||||||
#ordered.v
|
#ordered.v
|
||||||
|
|
|
@ -1,46 +0,0 @@
|
||||||
Require Import HoTT HitTactics.
|
|
||||||
Require Import lattice representations.definition fsets.operations extensionality.
|
|
||||||
|
|
||||||
Definition Sub A := A -> hProp.
|
|
||||||
|
|
||||||
Section k_finite.
|
|
||||||
|
|
||||||
Context {A : Type}.
|
|
||||||
Context `{Univalence}.
|
|
||||||
|
|
||||||
Instance subA_set : IsHSet (Sub A).
|
|
||||||
Proof.
|
|
||||||
apply _.
|
|
||||||
Defined.
|
|
||||||
|
|
||||||
Definition map : FSet A -> Sub A.
|
|
||||||
Proof.
|
|
||||||
unfold Sub.
|
|
||||||
intros X a.
|
|
||||||
apply (isIn a X).
|
|
||||||
Defined.
|
|
||||||
|
|
||||||
Definition k_finite (B : Sub A) : hProp.
|
|
||||||
Proof.
|
|
||||||
simple refine (@BuildhProp (@sig (FSet A) (fun X => B = map X)) _).
|
|
||||||
apply hprop_allpath.
|
|
||||||
unfold map.
|
|
||||||
intros [X PX] [Y PY].
|
|
||||||
assert (X0 : (fun a : A => a ∈ X) = (fun a : A => a ∈ Y)).
|
|
||||||
{
|
|
||||||
transitivity B.
|
|
||||||
- symmetry ; apply PX.
|
|
||||||
- apply PY.
|
|
||||||
}
|
|
||||||
assert (X = Y).
|
|
||||||
{
|
|
||||||
apply fset_ext.
|
|
||||||
intro a.
|
|
||||||
refine (ap (fun f : A -> hProp => f a) X0).
|
|
||||||
}
|
|
||||||
apply path_sigma_uncurried.
|
|
||||||
exists X1.
|
|
||||||
apply set_path2.
|
|
||||||
Defined.
|
|
||||||
|
|
||||||
End k_finite.
|
|
|
@ -0,0 +1,54 @@
|
||||||
|
Require Import HoTT HitTactics.
|
||||||
|
Require Import lattice representations.definition fsets.operations extensionality.
|
||||||
|
|
||||||
|
Definition Sub A := A -> hProp.
|
||||||
|
|
||||||
|
Section k_finite.
|
||||||
|
|
||||||
|
Context {A : Type}.
|
||||||
|
Context `{Univalence}.
|
||||||
|
|
||||||
|
Instance subA_set : IsHSet (Sub A).
|
||||||
|
Proof.
|
||||||
|
apply _.
|
||||||
|
Defined.
|
||||||
|
|
||||||
|
Definition map (X : FSet A) : Sub A := fun a => isIn a X.
|
||||||
|
|
||||||
|
Instance map_injective : IsEmbedding map.
|
||||||
|
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).
|
||||||
|
|
||||||
|
Lemma Kf_unfold : Kf <-> (exists (X : FSet A), forall (a : A), map X a).
|
||||||
|
Proof.
|
||||||
|
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.
|
Loading…
Reference in New Issue