HITs-Examples/FiniteSets/fsets/operations_decidable.v

48 lines
1.2 KiB
Coq
Raw Normal View History

2017-08-01 15:12:59 +02:00
(* Operations on [FSet A] when [A] has decidable equality *)
Require Import HoTT HitTactics.
Require Export representations.definition fsets.operations.
2017-08-01 15:12:59 +02:00
Section decidable_A.
Context {A : Type}.
Context {A_deceq : DecidablePaths A}.
Context `{Univalence}.
2017-08-07 16:57:21 +02:00
Global Instance isIn_decidable : forall (a : A) (X : FSet A), Decidable (a X).
2017-08-01 15:12:59 +02:00
Proof.
intros a.
hinduction ; try (intros ; apply path_ishprop).
- apply _.
2017-08-03 12:21:34 +02:00
- intros.
unfold Decidable.
destruct (dec (a = a0)) as [p | np].
* apply (inl (tr p)).
* right.
intro p.
strip_truncations.
contradiction.
2017-08-01 15:12:59 +02:00
- intros. apply _.
Defined.
Definition isIn_b (a : A) (X : FSet A) :=
2017-08-07 16:57:21 +02:00
match dec (a X) with
2017-08-01 15:12:59 +02:00
| inl _ => true
| inr _ => false
end.
2017-08-07 16:57:21 +02:00
Global Instance subset_decidable : forall (X Y : FSet A), Decidable (X Y).
2017-08-01 15:12:59 +02:00
Proof.
hinduction ; try (intros ; apply path_ishprop).
- intro ; apply _.
- intros ; apply _.
- intros ; apply _.
Defined.
Definition subset_b (X Y : FSet A) :=
2017-08-07 16:57:21 +02:00
match dec (X Y) with
2017-08-01 15:12:59 +02:00
| inl _ => true
| inr _ => false
end.
Definition intersection (X Y : FSet A) : FSet A := comprehension (fun a => isIn_b a Y) X.
End decidable_A.