1
0
mirror of https://github.com/nmvdw/HITs-Examples synced 2025-12-15 14:43:51 +01:00

Split the development into different directories

This commit is contained in:
Niels
2017-08-01 15:41:53 +02:00
parent bae04a6d2b
commit 0de37d6cea
13 changed files with 30 additions and 19 deletions

View File

@@ -0,0 +1,41 @@
(* 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 (isIn a X).
Proof.
intros a.
hinduction ; try (intros ; apply path_ishprop).
- apply _.
- intros. apply _.
- intros. apply _.
Defined.
Definition isIn_b (a : A) (X : FSet A) :=
match dec (isIn a X) with
| inl _ => true
| inr _ => false
end.
Global Instance subset_decidable : forall (X Y : FSet A), Decidable (subset X Y).
Proof.
hinduction ; try (intros ; apply path_ishprop).
- intro ; apply _.
- intros ; apply _.
- intros ; apply _.
Defined.
Definition subset_b (X Y : FSet A) :=
match dec (subset X Y) with
| inl _ => true
| inr _ => false
end.
Definition intersection (X Y : FSet A) : FSet A := comprehension (fun a => isIn_b a Y) X.
End decidable_A.