mirror of
https://github.com/nmvdw/HITs-Examples
synced 2025-11-03 15:13:51 +01:00
A negligible change in the structure
This commit is contained in:
141
FiniteSets/kuratowski/extensionality.v
Normal file
141
FiniteSets/kuratowski/extensionality.v
Normal file
@@ -0,0 +1,141 @@
|
||||
(** Extensionality of the FSets *)
|
||||
Require Import HoTT HitTactics.
|
||||
Require Import kuratowski.kuratowski_sets.
|
||||
|
||||
Section ext.
|
||||
Context {A : Type}.
|
||||
Context `{Univalence}.
|
||||
|
||||
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.
|
||||
|
||||
Lemma subset_union (X Y : FSet A) :
|
||||
X ⊆ Y -> X ∪ Y = Y.
|
||||
Proof.
|
||||
hinduction X ; try (intros ; apply path_ishprop).
|
||||
- intros.
|
||||
apply nl.
|
||||
- intros a.
|
||||
hinduction Y ; try (intros ; apply path_ishprop).
|
||||
+ intro.
|
||||
contradiction.
|
||||
+ 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).
|
||||
reflexivity.
|
||||
++ rewrite comm, <- assoc, (comm X2), (IH2 t).
|
||||
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.
|
||||
hinduction X ; try (intros ; apply path_ishprop).
|
||||
- apply (fun _ => tt).
|
||||
- intros.
|
||||
apply (tr(inl(tr idpath))).
|
||||
- intros X1 X2 HX1 HX2 Y.
|
||||
split ; unfold subset in *.
|
||||
* rewrite <- assoc.
|
||||
apply HX1.
|
||||
* rewrite (comm X1 X2), <- assoc.
|
||||
apply HX2.
|
||||
Defined.
|
||||
|
||||
Lemma subset_union_equiv
|
||||
: forall X Y : FSet A, X ⊆ Y <~> X ∪ Y = Y.
|
||||
Proof.
|
||||
intros X Y.
|
||||
eapply equiv_iff_hprop_uncurried.
|
||||
split.
|
||||
- apply subset_union.
|
||||
- intro HXY.
|
||||
rewrite <- HXY.
|
||||
apply subset_union_l.
|
||||
Defined.
|
||||
|
||||
Lemma subset_isIn (X Y : FSet A) :
|
||||
X ⊆ Y <~> forall (a : A), a ∈ X -> a ∈ Y.
|
||||
Proof.
|
||||
etransitivity.
|
||||
- apply subset_union_equiv.
|
||||
- eapply equiv_iff_hprop_uncurried ; split.
|
||||
* apply equiv_subset1_l.
|
||||
* apply equiv_subset1_r.
|
||||
Defined.
|
||||
|
||||
Lemma eq_subset (X Y : FSet A) :
|
||||
X = Y <~> (Y ⊆ X * X ⊆ Y).
|
||||
Proof.
|
||||
etransitivity ((Y ∪ X = X) * (X ∪ Y = Y)).
|
||||
- apply eq_subset2.
|
||||
- symmetry.
|
||||
eapply equiv_functor_prod' ; apply subset_union_equiv.
|
||||
Defined.
|
||||
End ext.
|
||||
165
FiniteSets/kuratowski/kuratowski_sets.v
Normal file
165
FiniteSets/kuratowski/kuratowski_sets.v
Normal file
@@ -0,0 +1,165 @@
|
||||
(** Definitions of the Kuratowski-finite sets via a HIT *)
|
||||
Require Import HoTT HitTactics.
|
||||
Require Export set_names lattice_examples.
|
||||
|
||||
Module Export FSet.
|
||||
Section FSet.
|
||||
Private Inductive FSet (A : Type) : Type :=
|
||||
| E : FSet A
|
||||
| L : A -> FSet A
|
||||
| U : FSet A -> FSet A -> FSet A.
|
||||
|
||||
Global Instance fset_empty : forall A, hasEmpty (FSet A) := E.
|
||||
Global Instance fset_singleton : forall A, hasSingleton (FSet A) A := L.
|
||||
Global Instance fset_union : forall A, hasUnion (FSet A) := U.
|
||||
|
||||
Variable A : Type.
|
||||
|
||||
Axiom assoc : forall (x y z : FSet A),
|
||||
x ∪ (y ∪ z) = (x ∪ y) ∪ z.
|
||||
|
||||
Axiom comm : forall (x y : FSet A),
|
||||
x ∪ y = y ∪ x.
|
||||
|
||||
Axiom nl : forall (x : FSet A),
|
||||
∅ ∪ x = x.
|
||||
|
||||
Axiom nr : forall (x : FSet A),
|
||||
x ∪ ∅ = x.
|
||||
|
||||
Axiom idem : forall (x : A),
|
||||
{|x|} ∪ {|x|} = {|x|}.
|
||||
|
||||
Axiom trunc : IsHSet (FSet A).
|
||||
|
||||
End FSet.
|
||||
|
||||
Arguments assoc {_} _ _ _.
|
||||
Arguments comm {_} _ _.
|
||||
Arguments nl {_} _.
|
||||
Arguments nr {_} _.
|
||||
Arguments idem {_} _.
|
||||
|
||||
Section FSet_induction.
|
||||
Variable (A : Type)
|
||||
(P : FSet A -> Type)
|
||||
(H : forall X : FSet A, IsHSet (P X))
|
||||
(eP : P ∅)
|
||||
(lP : forall a: A, P {|a|})
|
||||
(uP : forall (x y: FSet A), P x -> P y -> P (x ∪ y))
|
||||
(assocP : forall (x y z : FSet A)
|
||||
(px: P x) (py: P y) (pz: P z),
|
||||
assoc x y z #
|
||||
(uP x (y ∪ z) px (uP y z py pz))
|
||||
=
|
||||
(uP (x ∪ y) z (uP x y px py) pz))
|
||||
(commP : forall (x y: FSet A) (px: P x) (py: P y),
|
||||
comm x y #
|
||||
uP x y px py = uP y x py px)
|
||||
(nlP : forall (x : FSet A) (px: P x),
|
||||
nl x # uP ∅ x eP px = px)
|
||||
(nrP : forall (x : FSet A) (px: P x),
|
||||
nr x # uP x ∅ px eP = px)
|
||||
(idemP : forall (x : A),
|
||||
idem x # uP {|x|} {|x|} (lP x) (lP x) = lP x).
|
||||
|
||||
(* Induction principle *)
|
||||
Fixpoint FSet_ind
|
||||
(x : FSet A)
|
||||
{struct x}
|
||||
: P x
|
||||
:= (match x return _ -> _ -> _ -> _ -> _ -> _ -> P x with
|
||||
| E => fun _ _ _ _ _ _ => eP
|
||||
| L a => fun _ _ _ _ _ _ => lP a
|
||||
| U y z => fun _ _ _ _ _ _ => uP y z (FSet_ind y) (FSet_ind z)
|
||||
end) H assocP commP nlP nrP idemP.
|
||||
End FSet_induction.
|
||||
|
||||
Section FSet_recursion.
|
||||
Variable (A : Type)
|
||||
(P : Type)
|
||||
(H: IsHSet P)
|
||||
(e : P)
|
||||
(l : A -> P)
|
||||
(u : P -> P -> P)
|
||||
(assocP : forall (x y z : P), u x (u y z) = u (u x y) z)
|
||||
(commP : forall (x y : P), u x y = u y x)
|
||||
(nlP : forall (x : P), u e x = x)
|
||||
(nrP : forall (x : P), u x e = x)
|
||||
(idemP : forall (x : A), u (l x) (l x) = l x).
|
||||
|
||||
(* Recursion princople *)
|
||||
Definition FSet_rec : FSet A -> P.
|
||||
Proof.
|
||||
simple refine (FSet_ind A _ _ _ _ _ _ _ _ _ _)
|
||||
; try (intros ; simple refine ((transport_const _ _) @ _)) ; cbn.
|
||||
- apply e.
|
||||
- apply l.
|
||||
- intros x y ; apply u.
|
||||
- apply assocP.
|
||||
- apply commP.
|
||||
- apply nlP.
|
||||
- apply nrP.
|
||||
- apply idemP.
|
||||
Defined.
|
||||
End FSet_recursion.
|
||||
|
||||
Instance FSet_recursion A : HitRecursion (FSet A) :=
|
||||
{
|
||||
indTy := _; recTy := _;
|
||||
H_inductor := FSet_ind A; H_recursor := FSet_rec A
|
||||
}.
|
||||
End FSet.
|
||||
|
||||
Lemma union_idem {A : Type} : forall x: FSet A, x ∪ x = x.
|
||||
Proof.
|
||||
hinduction ; try (intros ; apply set_path2).
|
||||
- apply nl.
|
||||
- apply idem.
|
||||
- intros x y P Q.
|
||||
rewrite assoc.
|
||||
rewrite (comm x y).
|
||||
rewrite <- (assoc y x x).
|
||||
rewrite P.
|
||||
rewrite (comm y x).
|
||||
rewrite <- (assoc x y y).
|
||||
apply (ap (x ∪) Q).
|
||||
Defined.
|
||||
|
||||
Section relations.
|
||||
Context `{Univalence}.
|
||||
|
||||
(** Membership of finite sets. *)
|
||||
Global Instance fset_member : forall A, hasMembership (FSet A) A.
|
||||
Proof.
|
||||
intros A a.
|
||||
hrecursion.
|
||||
- apply False_hp.
|
||||
- apply (fun a' => merely(a = a')).
|
||||
- apply lor.
|
||||
- eauto with lattice_hints typeclass_instances.
|
||||
- eauto with lattice_hints typeclass_instances.
|
||||
- eauto with lattice_hints typeclass_instances.
|
||||
- eauto with lattice_hints typeclass_instances.
|
||||
- eauto with lattice_hints typeclass_instances.
|
||||
Defined.
|
||||
|
||||
(** Subset relation of finite sets. *)
|
||||
Global Instance fset_subset : forall A, hasSubset (FSet A).
|
||||
Proof.
|
||||
intros A X Y.
|
||||
hrecursion X.
|
||||
- apply Unit_hp.
|
||||
- apply (fun a => a ∈ Y).
|
||||
- intros X1 X2.
|
||||
exists (prod X1 X2).
|
||||
exact _.
|
||||
- eauto with lattice_hints typeclass_instances.
|
||||
- eauto with lattice_hints typeclass_instances.
|
||||
- intros.
|
||||
apply path_trunctype ; apply prod_unit_l.
|
||||
- intros.
|
||||
apply path_trunctype ; apply prod_unit_r.
|
||||
- eauto with lattice_hints typeclass_instances.
|
||||
Defined.
|
||||
End relations.
|
||||
84
FiniteSets/kuratowski/monad.v
Normal file
84
FiniteSets/kuratowski/monad.v
Normal file
@@ -0,0 +1,84 @@
|
||||
(* [FSet] is a (strong and stable) finite powerset monad *)
|
||||
Require Import HoTT HitTactics.
|
||||
Require Export representations.definition fsets.properties fsets.operations.
|
||||
|
||||
Definition ffmap {A B : Type} : (A -> B) -> FSet A -> FSet B.
|
||||
Proof.
|
||||
intro f.
|
||||
hrecursion.
|
||||
- exact ∅.
|
||||
- intro a. exact {| f a |}.
|
||||
- intros X Y. apply (X ∪ Y).
|
||||
- apply assoc.
|
||||
- apply comm.
|
||||
- apply nl.
|
||||
- apply nr.
|
||||
- simpl. intro x. apply idem.
|
||||
Defined.
|
||||
|
||||
Lemma ffmap_1 `{Funext} {A : Type} : @ffmap A A idmap = idmap.
|
||||
Proof.
|
||||
apply path_forall.
|
||||
intro x. hinduction x; try (intros; f_ap);
|
||||
try (intros; apply set_path2).
|
||||
Defined.
|
||||
|
||||
Global Instance fset_functorish `{Funext}: Functorish FSet
|
||||
:= { fmap := @ffmap; fmap_idmap := @ffmap_1 _ }.
|
||||
|
||||
Lemma ffmap_compose {A B C : Type} `{Funext} (f : A -> B) (g : B -> C) :
|
||||
fmap FSet (g o f) = fmap _ g o fmap _ f.
|
||||
Proof.
|
||||
apply path_forall. intro x.
|
||||
hrecursion x; try (intros; f_ap);
|
||||
try (intros; apply set_path2).
|
||||
Defined.
|
||||
|
||||
Definition join {A : Type} : FSet (FSet A) -> FSet A.
|
||||
Proof.
|
||||
hrecursion.
|
||||
- exact ∅.
|
||||
- exact idmap.
|
||||
- intros X Y. apply (X ∪ Y).
|
||||
- apply assoc.
|
||||
- apply comm.
|
||||
- apply nl.
|
||||
- apply nr.
|
||||
- simpl. apply union_idem.
|
||||
Defined.
|
||||
|
||||
Lemma join_assoc {A : Type} (X : FSet (FSet (FSet A))) :
|
||||
join (ffmap join X) = join (join X).
|
||||
Proof.
|
||||
hrecursion X; try (intros; f_ap);
|
||||
try (intros; apply set_path2).
|
||||
Defined.
|
||||
|
||||
Lemma join_return_1 {A : Type} (X : FSet A) :
|
||||
join ({| X |}) = X.
|
||||
Proof. reflexivity. Defined.
|
||||
|
||||
Lemma join_return_fmap {A : Type} (X : FSet A) :
|
||||
join ({| X |}) = join (ffmap (fun x => {|x|}) X).
|
||||
Proof.
|
||||
hrecursion X; try (intros; f_ap);
|
||||
try (intros; apply set_path2).
|
||||
Defined.
|
||||
|
||||
Lemma join_fmap_return_1 {A : Type} (X : FSet A) :
|
||||
join (ffmap (fun x => {|x|}) X) = X.
|
||||
Proof. refine ((join_return_fmap _)^ @ join_return_1 _). Defined.
|
||||
|
||||
Lemma fmap_isIn `{Univalence} {A B : Type} (f : A -> B) (a : A) (X : FSet A) :
|
||||
a ∈ X -> (f a) ∈ (ffmap f X).
|
||||
Proof.
|
||||
hinduction X; try (intros; apply path_ishprop).
|
||||
- apply idmap.
|
||||
- intros b Hab; strip_truncations.
|
||||
apply (tr (ap f Hab)).
|
||||
- intros X0 X1 HX0 HX1 Ha.
|
||||
strip_truncations. apply tr.
|
||||
destruct Ha as [Ha | Ha].
|
||||
+ left. by apply HX0.
|
||||
+ right. by apply HX1.
|
||||
Defined.
|
||||
183
FiniteSets/kuratowski/operations.v
Normal file
183
FiniteSets/kuratowski/operations.v
Normal file
@@ -0,0 +1,183 @@
|
||||
(** Operations on the [FSet A] for an arbitrary [A] *)
|
||||
Require Import HoTT HitTactics.
|
||||
Require Import kuratowski_sets monad_interface.
|
||||
|
||||
Section operations.
|
||||
(** 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.
|
||||
|
||||
Global Instance fset_pure : hasPure FSet.
|
||||
Proof.
|
||||
split.
|
||||
apply (fun _ a => {|a|}).
|
||||
Defined.
|
||||
|
||||
Global Instance fset_bind : hasBind FSet.
|
||||
Proof.
|
||||
split.
|
||||
intros A.
|
||||
hrecursion.
|
||||
- exact ∅.
|
||||
- exact idmap.
|
||||
- apply (∪).
|
||||
- apply assoc.
|
||||
- apply comm.
|
||||
- apply nl.
|
||||
- apply nr.
|
||||
- apply union_idem.
|
||||
Defined.
|
||||
|
||||
(** Set-theoretical operations. *)
|
||||
Global Instance fset_comprehension : forall A, hasComprehension (FSet A) A.
|
||||
Proof.
|
||||
intros A P.
|
||||
hrecursion.
|
||||
- apply ∅.
|
||||
- intro a.
|
||||
refine (if (P a) then {|a|} else ∅).
|
||||
- apply (∪).
|
||||
- apply assoc.
|
||||
- apply comm.
|
||||
- apply nl.
|
||||
- apply nr.
|
||||
- intros; simpl.
|
||||
destruct (P x).
|
||||
+ apply idem.
|
||||
+ apply nl.
|
||||
Defined.
|
||||
|
||||
Definition isEmpty (A : Type) :
|
||||
FSet A -> Bool.
|
||||
Proof.
|
||||
simple refine (FSet_rec _ _ _ true (fun _ => false) andb _ _ _ _ _)
|
||||
; try eauto with bool_lattice_hints typeclass_instances.
|
||||
Defined.
|
||||
|
||||
Definition single_product {A B : Type} (a : A) : FSet B -> FSet (A * B).
|
||||
Proof.
|
||||
hrecursion.
|
||||
- apply ∅.
|
||||
- intro b.
|
||||
apply {|(a, b)|}.
|
||||
- apply (∪).
|
||||
- apply assoc.
|
||||
- apply comm.
|
||||
- apply nl.
|
||||
- apply nr.
|
||||
- intros.
|
||||
apply idem.
|
||||
Defined.
|
||||
|
||||
Definition product {A B : Type} : FSet A -> FSet B -> FSet (A * B).
|
||||
Proof.
|
||||
intros X Y.
|
||||
hrecursion X.
|
||||
- apply ∅.
|
||||
- intro a.
|
||||
apply (single_product a Y).
|
||||
- apply (∪).
|
||||
- apply assoc.
|
||||
- apply comm.
|
||||
- apply nl.
|
||||
- apply nr.
|
||||
- intros ; apply union_idem.
|
||||
Defined.
|
||||
|
||||
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).
|
||||
|
||||
Context `{Univalence}.
|
||||
|
||||
(** Separation axiom for finite sets. *)
|
||||
Definition separation (A B : Type) : forall (X : FSet A) (f : {a | a ∈ X} -> B), FSet B.
|
||||
Proof.
|
||||
hinduction.
|
||||
- intros f.
|
||||
apply ∅.
|
||||
- intros a f.
|
||||
apply {|f (a; tr idpath)|}.
|
||||
- intros X1 X2 HX1 HX2 f.
|
||||
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)))).
|
||||
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.
|
||||
Defined.
|
||||
End operations.
|
||||
|
||||
Section operations_decidable.
|
||||
Context {A : Type}.
|
||||
Context {A_deceq : DecidablePaths A}.
|
||||
Context `{Univalence}.
|
||||
|
||||
Global Instance isIn_decidable (a : A) : forall (X : FSet A),
|
||||
Decidable (a ∈ X).
|
||||
Proof.
|
||||
hinduction ; try (intros ; apply path_ishprop).
|
||||
- apply _.
|
||||
- intros b.
|
||||
destruct (dec (a = b)) as [p | np].
|
||||
* apply (inl (tr p)).
|
||||
* refine (inr(fun p => _)).
|
||||
strip_truncations.
|
||||
contradiction.
|
||||
- apply _.
|
||||
Defined.
|
||||
|
||||
Global Instance fset_member_bool : hasMembership_decidable (FSet A) A.
|
||||
Proof.
|
||||
intros a X.
|
||||
refine (if (dec a ∈ X) then true else false).
|
||||
Defined.
|
||||
|
||||
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)
|
||||
:= fun X Y => {|X & (fun a => a ∈_d Y)|}.
|
||||
End operations_decidable.
|
||||
54
FiniteSets/kuratowski/operations_decidable.v
Normal file
54
FiniteSets/kuratowski/operations_decidable.v
Normal file
@@ -0,0 +1,54 @@
|
||||
(* Operations on [FSet A] when [A] has decidable equality *)
|
||||
Require Import HoTT HitTactics.
|
||||
Require Export representations.definition fsets.operations.
|
||||
|
||||
Section decidable_A.
|
||||
Context {A : Type}.
|
||||
Context {A_deceq : DecidablePaths A}.
|
||||
Context `{Univalence}.
|
||||
|
||||
Global Instance isIn_decidable : forall (a : A) (X : FSet A),
|
||||
Decidable (a ∈ X).
|
||||
Proof.
|
||||
intros a.
|
||||
hinduction ; try (intros ; apply path_ishprop).
|
||||
- apply _.
|
||||
- intros.
|
||||
unfold Decidable.
|
||||
destruct (dec (a = a0)) as [p | np].
|
||||
* apply (inl (tr p)).
|
||||
* right.
|
||||
intro p.
|
||||
strip_truncations.
|
||||
contradiction.
|
||||
- intros. apply _.
|
||||
Defined.
|
||||
|
||||
Global Instance fset_member_bool : hasMembership_decidable (FSet A) A.
|
||||
Proof.
|
||||
intros a X.
|
||||
destruct (dec (a ∈ X)).
|
||||
- apply true.
|
||||
- apply false.
|
||||
Defined.
|
||||
|
||||
Global Instance subset_decidable : forall (X Y : FSet A), Decidable (X ⊆ Y).
|
||||
Proof.
|
||||
hinduction ; try (intros ; apply path_ishprop).
|
||||
- intro ; apply _.
|
||||
- intros ; apply _.
|
||||
- intros. unfold subset in *. apply _.
|
||||
Defined.
|
||||
|
||||
Global Instance fset_subset_bool : hasSubset_decidable (FSet A).
|
||||
Proof.
|
||||
intros X Y.
|
||||
destruct (dec (X ⊆ Y)).
|
||||
- apply true.
|
||||
- apply false.
|
||||
Defined.
|
||||
|
||||
Global Instance fset_intersection : hasIntersection (FSet A)
|
||||
:= fun X Y => {|X & (fun a => a ∈_d Y)|}.
|
||||
|
||||
End decidable_A.
|
||||
529
FiniteSets/kuratowski/properties.v
Normal file
529
FiniteSets/kuratowski/properties.v
Normal file
@@ -0,0 +1,529 @@
|
||||
Require Import HoTT HitTactics.
|
||||
Require Import kuratowski.extensionality kuratowski.operations kuratowski_sets.
|
||||
Require Import lattice_interface lattice_examples monad_interface.
|
||||
|
||||
(** Lemmas relating operations to the membership predicate *)
|
||||
Section properties_membership.
|
||||
Context {A : Type} `{Univalence}.
|
||||
|
||||
Definition empty_isIn (a: A) : a ∈ ∅ -> Empty := idmap.
|
||||
|
||||
Definition singleton_isIn (a b: A) : a ∈ {|b|} -> Trunc (-1) (a = b) := idmap.
|
||||
|
||||
Definition union_isIn (X Y : FSet A) (a : A)
|
||||
: a ∈ X ∪ Y = a ∈ X ∨ a ∈ Y := idpath.
|
||||
|
||||
Lemma comprehension_union (ϕ : A -> Bool) : forall X Y : FSet A,
|
||||
{|X ∪ Y & ϕ|} = {|X & ϕ|} ∪ {|Y & ϕ|}.
|
||||
Proof.
|
||||
reflexivity.
|
||||
Defined.
|
||||
|
||||
Lemma comprehension_isIn (ϕ : A -> Bool) (a : A) : forall X : FSet A,
|
||||
a ∈ {|X & ϕ|} = if ϕ a then a ∈ X else False_hp.
|
||||
Proof.
|
||||
hinduction ; try (intros ; apply set_path2).
|
||||
- destruct (ϕ a) ; reflexivity.
|
||||
- intros b.
|
||||
assert (forall c d, ϕ a = c -> ϕ b = d ->
|
||||
a ∈ (if ϕ b then {|b|} else ∅)
|
||||
=
|
||||
(if ϕ a then BuildhProp (Trunc (-1) (a = b)) else False_hp)) as X.
|
||||
{
|
||||
intros c d Hc Hd.
|
||||
destruct c ; destruct d ; rewrite Hc, Hd ; try reflexivity
|
||||
; apply path_iff_hprop ; try contradiction ; intros ; strip_truncations
|
||||
; apply (false_ne_true).
|
||||
* apply (Hd^ @ ap ϕ X^ @ Hc).
|
||||
* apply (Hc^ @ ap ϕ X @ Hd).
|
||||
}
|
||||
apply (X (ϕ a) (ϕ b) idpath idpath).
|
||||
- intros X Y H1 H2.
|
||||
rewrite comprehension_union.
|
||||
rewrite union_isIn.
|
||||
rewrite H1, H2.
|
||||
destruct (ϕ a).
|
||||
* reflexivity.
|
||||
* apply path_iff_hprop.
|
||||
** intros Z ; strip_truncations.
|
||||
destruct Z ; assumption.
|
||||
** intros ; apply tr ; right ; assumption.
|
||||
Defined.
|
||||
|
||||
Context {B : Type}.
|
||||
|
||||
Lemma isIn_singleproduct (a : A) (b : B) (c : A) : forall (Y : FSet B),
|
||||
(a, b) ∈ (single_product c Y) = land (BuildhProp (Trunc (-1) (a = c))) (b ∈ Y).
|
||||
Proof.
|
||||
hinduction ; try (intros ; apply path_ishprop).
|
||||
- apply path_hprop ; symmetry ; apply prod_empty_r.
|
||||
- intros d.
|
||||
apply path_iff_hprop.
|
||||
* intros.
|
||||
strip_truncations.
|
||||
split ; apply tr ; try (apply (ap fst X)) ; try (apply (ap snd X)).
|
||||
* intros [Z1 Z2].
|
||||
strip_truncations.
|
||||
rewrite Z1, Z2.
|
||||
apply (tr idpath).
|
||||
- intros X1 X2 HX1 HX2.
|
||||
apply path_iff_hprop ; rewrite ?union_isIn.
|
||||
* intros X.
|
||||
cbn in *.
|
||||
strip_truncations.
|
||||
destruct X as [H1 | H1] ; rewrite ?HX1, ?HX2 in H1
|
||||
; destruct H1 as [H1 H2].
|
||||
** split.
|
||||
*** apply H1.
|
||||
*** apply (tr(inl H2)).
|
||||
** split.
|
||||
*** apply H1.
|
||||
*** apply (tr(inr H2)).
|
||||
* intros [H1 H2].
|
||||
cbn in *.
|
||||
strip_truncations.
|
||||
apply tr.
|
||||
rewrite HX1, HX2.
|
||||
destruct H2 as [Hb1 | Hb2].
|
||||
** left.
|
||||
split ; try (apply (tr H1)) ; try (apply Hb1).
|
||||
** right.
|
||||
split ; try (apply (tr H1)) ; try (apply Hb2).
|
||||
Defined.
|
||||
|
||||
Definition isIn_product (a : A) (b : B) (X : FSet A) (Y : FSet B) :
|
||||
(a,b) ∈ (product X Y) = land (a ∈ X) (b ∈ Y).
|
||||
Proof.
|
||||
hinduction X ; try (intros ; apply path_ishprop).
|
||||
- apply path_hprop ; symmetry ; apply prod_empty_l.
|
||||
- intros.
|
||||
apply isIn_singleproduct.
|
||||
- intros X1 X2 HX1 HX2.
|
||||
cbn.
|
||||
rewrite HX1, HX2.
|
||||
apply path_iff_hprop ; rewrite ?union_isIn.
|
||||
* intros X.
|
||||
strip_truncations.
|
||||
destruct X as [[H3 H4] | [H3 H4]] ; split ; try (apply H4).
|
||||
** apply (tr(inl H3)).
|
||||
** apply (tr(inr H3)).
|
||||
* intros [H1 H2].
|
||||
strip_truncations.
|
||||
destruct H1 as [H1 | H1] ; apply tr.
|
||||
** left ; split ; assumption.
|
||||
** right ; split ; assumption.
|
||||
Defined.
|
||||
|
||||
Lemma separation_isIn : forall (X : FSet A) (f : {a | a ∈ X} -> B) (b : B),
|
||||
b ∈ (separation A B X f) = hexists (fun a : A => hexists (fun (p : a ∈ X) => f (a;p) = b)).
|
||||
Proof.
|
||||
hinduction ; try (intros ; apply path_forall ; intro ; apply path_ishprop).
|
||||
- intros ; simpl.
|
||||
apply path_iff_hprop ; try contradiction.
|
||||
intros X.
|
||||
strip_truncations.
|
||||
destruct X as [a X].
|
||||
strip_truncations.
|
||||
destruct X as [p X].
|
||||
assumption.
|
||||
- intros.
|
||||
apply path_iff_hprop ; simpl.
|
||||
* intros ; strip_truncations.
|
||||
apply tr.
|
||||
exists a.
|
||||
apply tr.
|
||||
exists (tr idpath).
|
||||
apply X^.
|
||||
* intros X ; strip_truncations.
|
||||
destruct X as [a0 X].
|
||||
strip_truncations.
|
||||
destruct X as [X q].
|
||||
simple refine (Trunc_ind _ _ X).
|
||||
intros p.
|
||||
apply tr.
|
||||
rewrite <- q.
|
||||
f_ap.
|
||||
simple refine (path_sigma _ _ _ _ _).
|
||||
** apply p.
|
||||
** apply path_ishprop.
|
||||
- intros X1 X2 HX1 HX2 f b.
|
||||
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)))).
|
||||
specialize (HX1 fX1 b).
|
||||
specialize (HX2 fX2 b).
|
||||
apply path_iff_hprop.
|
||||
* intros X.
|
||||
cbn in *.
|
||||
strip_truncations.
|
||||
destruct X as [X | X].
|
||||
** rewrite HX1 in X.
|
||||
strip_truncations.
|
||||
destruct X as [a Ha].
|
||||
apply tr.
|
||||
exists a.
|
||||
strip_truncations.
|
||||
destruct Ha as [p pa].
|
||||
apply tr.
|
||||
exists (tr(inl p)).
|
||||
rewrite <- pa.
|
||||
reflexivity.
|
||||
** rewrite HX2 in X.
|
||||
strip_truncations.
|
||||
destruct X as [a Ha].
|
||||
apply tr.
|
||||
exists a.
|
||||
strip_truncations.
|
||||
destruct Ha as [p pa].
|
||||
apply tr.
|
||||
exists (tr(inr p)).
|
||||
rewrite <- pa.
|
||||
reflexivity.
|
||||
* intros.
|
||||
strip_truncations.
|
||||
destruct X as [a X].
|
||||
strip_truncations.
|
||||
destruct X as [Ha p].
|
||||
cbn.
|
||||
simple refine (Trunc_ind _ _ Ha) ; intros [Ha1 | Ha2].
|
||||
** refine (tr(inl _)).
|
||||
rewrite HX1.
|
||||
apply tr.
|
||||
exists a.
|
||||
apply tr.
|
||||
exists Ha1.
|
||||
rewrite <- p.
|
||||
unfold fX1.
|
||||
repeat f_ap.
|
||||
apply path_ishprop.
|
||||
** refine (tr(inr _)).
|
||||
rewrite HX2.
|
||||
apply tr.
|
||||
exists a.
|
||||
apply tr.
|
||||
exists Ha2.
|
||||
rewrite <- p.
|
||||
unfold fX2.
|
||||
repeat f_ap.
|
||||
apply path_ishprop.
|
||||
Defined.
|
||||
End properties_membership.
|
||||
|
||||
Ltac simplify_isIn :=
|
||||
repeat (rewrite union_isIn
|
||||
|| rewrite comprehension_isIn).
|
||||
|
||||
Ltac toHProp :=
|
||||
repeat intro;
|
||||
apply fset_ext ; intros ;
|
||||
simplify_isIn ; eauto with lattice_hints typeclass_instances.
|
||||
|
||||
(** [FSet A] is a join semilattice. *)
|
||||
Section fset_join_semilattice.
|
||||
Context {A : Type}.
|
||||
|
||||
Instance: bottom(FSet A).
|
||||
Proof.
|
||||
unfold bottom.
|
||||
apply ∅.
|
||||
Defined.
|
||||
|
||||
Instance: maximum(FSet A).
|
||||
Proof.
|
||||
intros x y.
|
||||
apply (x ∪ y).
|
||||
Defined.
|
||||
|
||||
Global Instance joinsemilattice_fset : JoinSemiLattice (FSet A).
|
||||
Proof.
|
||||
split.
|
||||
- intros ? ?.
|
||||
apply comm.
|
||||
- intros ? ? ?.
|
||||
apply (assoc _ _ _)^.
|
||||
- intros ?.
|
||||
apply union_idem.
|
||||
- intros x.
|
||||
apply nl.
|
||||
- intros ?.
|
||||
apply nr.
|
||||
Defined.
|
||||
End fset_join_semilattice.
|
||||
|
||||
(* Lemmas relating operations to the membership predicate *)
|
||||
Section properties_membership_decidable.
|
||||
Context {A : Type} `{DecidablePaths A} `{Univalence}.
|
||||
|
||||
Lemma ext : forall (S T : FSet A), (forall a, a ∈_d S = a ∈_d T) -> S = T.
|
||||
Proof.
|
||||
intros X Y H2.
|
||||
apply fset_ext.
|
||||
intro a.
|
||||
specialize (H2 a).
|
||||
unfold member_dec, fset_member_bool, dec in H2.
|
||||
destruct (isIn_decidable a X) ; destruct (isIn_decidable a Y).
|
||||
- apply path_iff_hprop ; intro ; assumption.
|
||||
- contradiction (true_ne_false).
|
||||
- contradiction (true_ne_false) ; apply H2^.
|
||||
- apply path_iff_hprop ; intro ; contradiction.
|
||||
Defined.
|
||||
|
||||
Lemma empty_isIn_d (a : A) :
|
||||
a ∈_d ∅ = false.
|
||||
Proof.
|
||||
reflexivity.
|
||||
Defined.
|
||||
|
||||
Lemma singleton_isIn_d (a b : A) :
|
||||
a ∈ {|b|} -> a = b.
|
||||
Proof.
|
||||
intros. strip_truncations. assumption.
|
||||
Defined.
|
||||
|
||||
Lemma singleton_isIn_d_true (a b : A) (p : a = b) :
|
||||
a ∈_d {|b|} = true.
|
||||
Proof.
|
||||
unfold member_dec, fset_member_bool, dec.
|
||||
destruct (isIn_decidable a {|b|}) as [n | n] .
|
||||
- reflexivity.
|
||||
- simpl in n.
|
||||
contradiction (n (tr p)).
|
||||
Defined.
|
||||
|
||||
Lemma singleton_isIn_d_aa (a : A) :
|
||||
a ∈_d {|a|} = true.
|
||||
Proof.
|
||||
apply singleton_isIn_d_true ; reflexivity.
|
||||
Defined.
|
||||
|
||||
Lemma singleton_isIn_d_false (a b : A) (p : a <> b) :
|
||||
a ∈_d {|b|} = false.
|
||||
Proof.
|
||||
unfold member_dec, fset_member_bool, dec in *.
|
||||
destruct (isIn_decidable a {|b|}).
|
||||
- simpl in t.
|
||||
strip_truncations.
|
||||
contradiction.
|
||||
- reflexivity.
|
||||
Defined.
|
||||
|
||||
Lemma union_isIn_d (X Y : FSet A) (a : A) :
|
||||
a ∈_d (X ∪ Y) = orb (a ∈_d X) (a ∈_d Y).
|
||||
Proof.
|
||||
unfold member_dec, fset_member_bool, dec.
|
||||
simpl.
|
||||
destruct (isIn_decidable a X) ; destruct (isIn_decidable a Y) ; reflexivity.
|
||||
Defined.
|
||||
|
||||
Lemma comprehension_isIn_d (Y : FSet A) (ϕ : A -> Bool) (a : A) :
|
||||
a ∈_d {|Y & ϕ|} = andb (a ∈_d Y) (ϕ a).
|
||||
Proof.
|
||||
unfold member_dec, fset_member_bool, dec ; simpl.
|
||||
destruct (isIn_decidable a {|Y & ϕ|}) as [t | t]
|
||||
; destruct (isIn_decidable a Y) as [n | n] ; rewrite comprehension_isIn in t
|
||||
; destruct (ϕ a) ; try reflexivity ; try contradiction
|
||||
; try (contradiction (n t)) ; contradiction (t n).
|
||||
Defined.
|
||||
|
||||
Lemma intersection_isIn_d (X Y: FSet A) (a : A) :
|
||||
a ∈_d (intersection X Y) = andb (a ∈_d X) (a ∈_d Y).
|
||||
Proof.
|
||||
apply comprehension_isIn_d.
|
||||
Defined.
|
||||
End properties_membership_decidable.
|
||||
|
||||
(* Some suporting tactics *)
|
||||
Ltac simplify_isIn_d :=
|
||||
repeat (rewrite union_isIn_d
|
||||
|| rewrite singleton_isIn_d_aa
|
||||
|| rewrite intersection_isIn_d
|
||||
|| rewrite comprehension_isIn_d).
|
||||
|
||||
Ltac toBool :=
|
||||
repeat intro;
|
||||
apply ext ; intros ;
|
||||
simplify_isIn_d ; eauto with bool_lattice_hints typeclass_instances.
|
||||
|
||||
(** If `A` has decidable equality, then `FSet A` is a lattice *)
|
||||
Section set_lattice.
|
||||
Context {A : Type}.
|
||||
Context {A_deceq : DecidablePaths A}.
|
||||
Context `{Univalence}.
|
||||
|
||||
Instance fset_max : maximum (FSet A).
|
||||
Proof.
|
||||
intros x y.
|
||||
apply (x ∪ y).
|
||||
Defined.
|
||||
|
||||
Instance fset_min : minimum (FSet A).
|
||||
Proof.
|
||||
intros x y.
|
||||
apply (x ∩ y).
|
||||
Defined.
|
||||
|
||||
Instance fset_bot : bottom (FSet A).
|
||||
Proof.
|
||||
unfold bottom.
|
||||
apply ∅.
|
||||
Defined.
|
||||
|
||||
Instance lattice_fset : Lattice (FSet A).
|
||||
Proof.
|
||||
split ; toBool.
|
||||
Defined.
|
||||
End set_lattice.
|
||||
|
||||
(** If `A` has decidable equality, then so has `FSet A`. *)
|
||||
Section dec_eq.
|
||||
Context {A : Type} `{DecidablePaths A} `{Univalence}.
|
||||
|
||||
Instance fsets_dec_eq : DecidablePaths (FSet A).
|
||||
Proof.
|
||||
intros X Y.
|
||||
apply (decidable_equiv' ((Y ⊆ X) * (X ⊆ Y)) (eq_subset X Y)^-1).
|
||||
apply decidable_prod.
|
||||
Defined.
|
||||
End dec_eq.
|
||||
|
||||
(** [FSet] is a (strong and stable) finite powerset monad *)
|
||||
Section monad_fset.
|
||||
Context `{Funext}.
|
||||
|
||||
Global Instance fset_functorish : Functorish FSet.
|
||||
Proof.
|
||||
simple refine (Build_Functorish _ _ _).
|
||||
- intros ? ? f.
|
||||
apply (fset_fmap f).
|
||||
- intros A.
|
||||
apply path_forall.
|
||||
intro x.
|
||||
hinduction x
|
||||
; try (intros ; f_ap)
|
||||
; try (intros ; apply path_ishprop).
|
||||
Defined.
|
||||
|
||||
Global Instance fset_functor : Functor FSet.
|
||||
Proof.
|
||||
split.
|
||||
intros.
|
||||
apply path_forall.
|
||||
intro x.
|
||||
hrecursion x
|
||||
; try (intros ; f_ap)
|
||||
; try (intros ; apply path_ishprop).
|
||||
Defined.
|
||||
|
||||
Global Instance fset_monad : Monad FSet.
|
||||
Proof.
|
||||
split.
|
||||
- intros.
|
||||
apply path_forall.
|
||||
intro X.
|
||||
hrecursion X ; try (intros; f_ap) ;
|
||||
try (intros; apply path_ishprop).
|
||||
- intros.
|
||||
apply path_forall.
|
||||
intro X.
|
||||
hrecursion X ; try (intros; f_ap) ;
|
||||
try (intros; apply path_ishprop).
|
||||
- intros.
|
||||
apply path_forall.
|
||||
intro X.
|
||||
hrecursion X ; try (intros; f_ap) ;
|
||||
try (intros; apply path_ishprop).
|
||||
Defined.
|
||||
|
||||
Lemma fmap_isIn `{Univalence} {A B : Type} (f : A -> B) (a : A) (X : FSet A) :
|
||||
a ∈ X -> (f a) ∈ (fmap FSet f X).
|
||||
Proof.
|
||||
hinduction X; try (intros; apply path_ishprop).
|
||||
- apply idmap.
|
||||
- intros b Hab; strip_truncations.
|
||||
apply (tr (ap f Hab)).
|
||||
- intros X0 X1 HX0 HX1 Ha.
|
||||
strip_truncations. apply tr.
|
||||
destruct Ha as [Ha | Ha].
|
||||
+ left. by apply HX0.
|
||||
+ right. by apply HX1.
|
||||
Defined.
|
||||
End monad_fset.
|
||||
|
||||
(** comprehension properties *)
|
||||
Section comprehension_properties.
|
||||
Context {A : Type} `{Univalence}.
|
||||
|
||||
Lemma comprehension_false : forall (X : FSet A), {|X & fun _ => false|} = ∅.
|
||||
Proof.
|
||||
toHProp.
|
||||
Defined.
|
||||
|
||||
Lemma comprehension_subset : forall ϕ (X : FSet A),
|
||||
{|X & ϕ|} ∪ X = X.
|
||||
Proof.
|
||||
toHProp.
|
||||
destruct (ϕ a) ; eauto with lattice_hints typeclass_instances.
|
||||
Defined.
|
||||
|
||||
Lemma comprehension_or : forall ϕ ψ (X : FSet A),
|
||||
{|X & (fun a => orb (ϕ a) (ψ a))|}
|
||||
= {|X & ϕ|} ∪ {|X & ψ|}.
|
||||
Proof.
|
||||
toHProp.
|
||||
symmetry ; destruct (ϕ a) ; destruct (ψ a)
|
||||
; eauto with lattice_hints typeclass_instances.
|
||||
Defined.
|
||||
|
||||
Lemma comprehension_all : forall (X : FSet A),
|
||||
{|X & fun _ => true|} = X.
|
||||
Proof.
|
||||
toHProp.
|
||||
Defined.
|
||||
End comprehension_properties.
|
||||
|
||||
(** For [FSet A] we have mere choice. *)
|
||||
Section mere_choice.
|
||||
Context {A : Type} `{Univalence}.
|
||||
|
||||
Local Ltac solve_mc :=
|
||||
repeat (let x := fresh in intro x ; try(destruct x))
|
||||
; simpl
|
||||
; rewrite transport_sum
|
||||
; try (f_ap ; apply path_ishprop)
|
||||
; try (f_ap ; apply set_path2).
|
||||
|
||||
Lemma merely_choice : forall X : FSet A, (X = ∅) + (hexists (fun a => a ∈ X)).
|
||||
Proof.
|
||||
hinduction.
|
||||
- apply (inl idpath).
|
||||
- intro a.
|
||||
refine (inr (tr (a ; tr idpath))).
|
||||
- intros X Y TX TY.
|
||||
destruct TX as [XE | HX] ; destruct TY as [YE | HY].
|
||||
* refine (inl _).
|
||||
rewrite XE, YE.
|
||||
apply (union_idem ∅).
|
||||
* right.
|
||||
strip_truncations.
|
||||
destruct HY as [a Ya].
|
||||
refine (tr _).
|
||||
exists a.
|
||||
apply (tr (inr Ya)).
|
||||
* right.
|
||||
strip_truncations.
|
||||
destruct HX as [a Xa].
|
||||
refine (tr _).
|
||||
exists a.
|
||||
apply (tr (inl Xa)).
|
||||
* right.
|
||||
strip_truncations.
|
||||
destruct (HX, HY) as [[a Xa] [b Yb]].
|
||||
refine (tr _).
|
||||
exists a.
|
||||
apply (tr (inl Xa)).
|
||||
- solve_mc.
|
||||
- solve_mc.
|
||||
- solve_mc.
|
||||
- solve_mc.
|
||||
- solve_mc.
|
||||
Defined.
|
||||
End mere_choice.
|
||||
145
FiniteSets/kuratowski/properties_decidable.v
Normal file
145
FiniteSets/kuratowski/properties_decidable.v
Normal file
@@ -0,0 +1,145 @@
|
||||
(* Properties of [FSet A] where [A] has decidable equality *)
|
||||
Require Import HoTT HitTactics.
|
||||
From fsets Require Export properties extensionality operations_decidable.
|
||||
Require Export lattice.
|
||||
|
||||
(* Lemmas relating operations to the membership predicate *)
|
||||
Section operations_isIn.
|
||||
|
||||
Context {A : Type} `{DecidablePaths A} `{Univalence}.
|
||||
|
||||
Lemma ext : forall (S T : FSet A), (forall a, a ∈_d S = a ∈_d T) -> S = T.
|
||||
Proof.
|
||||
intros X Y H2.
|
||||
apply fset_ext.
|
||||
intro a.
|
||||
specialize (H2 a).
|
||||
unfold member_dec, fset_member_bool, dec in H2.
|
||||
destruct (isIn_decidable a X) ; destruct (isIn_decidable a Y).
|
||||
- apply path_iff_hprop ; intro ; assumption.
|
||||
- contradiction (true_ne_false).
|
||||
- contradiction (true_ne_false) ; apply H2^.
|
||||
- apply path_iff_hprop ; intro ; contradiction.
|
||||
Defined.
|
||||
|
||||
Lemma empty_isIn (a : A) :
|
||||
a ∈_d ∅ = false.
|
||||
Proof.
|
||||
reflexivity.
|
||||
Defined.
|
||||
|
||||
Lemma L_isIn (a b : A) :
|
||||
a ∈ {|b|} -> a = b.
|
||||
Proof.
|
||||
intros. strip_truncations. assumption.
|
||||
Defined.
|
||||
|
||||
Lemma L_isIn_b_true (a b : A) (p : a = b) :
|
||||
a ∈_d {|b|} = true.
|
||||
Proof.
|
||||
unfold member_dec, fset_member_bool, dec.
|
||||
destruct (isIn_decidable a {|b|}) as [n | n] .
|
||||
- reflexivity.
|
||||
- simpl in n.
|
||||
contradiction (n (tr p)).
|
||||
Defined.
|
||||
|
||||
Lemma L_isIn_b_aa (a : A) :
|
||||
a ∈_d {|a|} = true.
|
||||
Proof.
|
||||
apply L_isIn_b_true ; reflexivity.
|
||||
Defined.
|
||||
|
||||
Lemma L_isIn_b_false (a b : A) (p : a <> b) :
|
||||
a ∈_d {|b|} = false.
|
||||
Proof.
|
||||
unfold member_dec, fset_member_bool, dec in *.
|
||||
destruct (isIn_decidable a {|b|}).
|
||||
- simpl in t.
|
||||
strip_truncations.
|
||||
contradiction.
|
||||
- reflexivity.
|
||||
Defined.
|
||||
|
||||
(* Union and membership *)
|
||||
Lemma union_isIn_b (X Y : FSet A) (a : A) :
|
||||
a ∈_d (X ∪ Y) = orb (a ∈_d X) (a ∈_d Y).
|
||||
Proof.
|
||||
unfold member_dec, fset_member_bool, dec.
|
||||
simpl.
|
||||
destruct (isIn_decidable a X) ; destruct (isIn_decidable a Y) ; reflexivity.
|
||||
Defined.
|
||||
|
||||
Lemma comprehension_isIn_b (Y : FSet A) (ϕ : A -> Bool) (a : A) :
|
||||
a ∈_d {|Y & ϕ|} = andb (a ∈_d Y) (ϕ a).
|
||||
Proof.
|
||||
unfold member_dec, fset_member_bool, dec ; simpl.
|
||||
destruct (isIn_decidable a {|Y & ϕ|}) as [t | t]
|
||||
; destruct (isIn_decidable a Y) as [n | n] ; rewrite comprehension_isIn in t
|
||||
; destruct (ϕ a) ; try reflexivity ; try contradiction
|
||||
; try (contradiction (n t)) ; contradiction (t n).
|
||||
Defined.
|
||||
|
||||
Lemma intersection_isIn_b (X Y: FSet A) (a : A) :
|
||||
a ∈_d (intersection X Y) = andb (a ∈_d X) (a ∈_d Y).
|
||||
Proof.
|
||||
apply comprehension_isIn_b.
|
||||
Defined.
|
||||
|
||||
End operations_isIn.
|
||||
|
||||
(* Some suporting tactics *)
|
||||
Ltac simplify_isIn_b :=
|
||||
repeat (rewrite union_isIn_b
|
||||
|| rewrite L_isIn_b_aa
|
||||
|| rewrite intersection_isIn_b
|
||||
|| rewrite comprehension_isIn_b).
|
||||
|
||||
Ltac toBool :=
|
||||
repeat intro;
|
||||
apply ext ; intros ;
|
||||
simplify_isIn_b ; eauto with bool_lattice_hints typeclass_instances.
|
||||
|
||||
Section SetLattice.
|
||||
|
||||
Context {A : Type}.
|
||||
Context {A_deceq : DecidablePaths A}.
|
||||
Context `{Univalence}.
|
||||
|
||||
Instance fset_max : maximum (FSet A).
|
||||
Proof.
|
||||
intros x y.
|
||||
apply (x ∪ y).
|
||||
Defined.
|
||||
|
||||
Instance fset_min : minimum (FSet A).
|
||||
Proof.
|
||||
intros x y.
|
||||
apply (x ∩ y).
|
||||
Defined.
|
||||
|
||||
Instance fset_bot : bottom (FSet A).
|
||||
Proof.
|
||||
unfold bottom.
|
||||
apply ∅.
|
||||
Defined.
|
||||
|
||||
Instance lattice_fset : Lattice (FSet A).
|
||||
Proof.
|
||||
split; toBool.
|
||||
Defined.
|
||||
|
||||
End SetLattice.
|
||||
|
||||
(* With extensionality we can prove decidable equality *)
|
||||
Section dec_eq.
|
||||
Context (A : Type) `{DecidablePaths A} `{Univalence}.
|
||||
|
||||
Instance fsets_dec_eq : DecidablePaths (FSet A).
|
||||
Proof.
|
||||
intros X Y.
|
||||
apply (decidable_equiv' ((Y ⊆ X) * (X ⊆ Y)) (eq_subset X Y)^-1).
|
||||
apply decidable_prod.
|
||||
Defined.
|
||||
|
||||
End dec_eq.
|
||||
Reference in New Issue
Block a user