HITs-Examples/FiniteSets/fsets/operations.v

105 lines
2.5 KiB
Coq
Raw Normal View History

2017-08-01 15:12:59 +02:00
(* Operations on the [FSet A] for an arbitrary [A] *)
Require Import HoTT HitTactics.
Require Import representations.definition disjunction lattice.
2017-05-23 16:30:31 +02:00
2017-08-01 15:12:59 +02:00
Section operations.
2017-08-07 16:49:46 +02:00
Context `{Univalence}.
2017-05-23 16:30:31 +02:00
2017-08-08 17:00:30 +02:00
Global Instance fset_member : forall A, hasMembership (FSet A) A.
2017-08-07 16:49:46 +02:00
Proof.
2017-08-08 17:00:30 +02:00
intros A a.
2017-08-07 16:49:46 +02:00
hrecursion.
- exists Empty.
exact _.
- intro a'.
exists (Trunc (-1) (a = a')).
exact _.
- apply lor.
- intros ; symmetry ; apply lor_assoc.
- apply lor_commutative.
- apply lor_nl.
- apply lor_nr.
- intros ; apply lor_idem.
Defined.
2017-05-23 16:30:31 +02:00
2017-08-08 17:00:30 +02:00
Global Instance fset_comprehension : forall A, hasComprehension (FSet A) A.
2017-08-07 16:49:46 +02:00
Proof.
2017-08-08 17:00:30 +02:00
intros A P.
2017-08-07 16:49:46 +02:00
hrecursion.
- apply .
- intro a.
refine (if (P a) then {|a|} else ).
2017-08-08 15:29:50 +02:00
- apply ().
2017-08-07 16:49:46 +02:00
- apply assoc.
- apply comm.
- apply nl.
- apply nr.
- intros; simpl.
destruct (P x).
+ apply idem.
+ apply nl.
Defined.
2017-05-23 16:30:31 +02:00
2017-08-08 17:00:30 +02:00
Definition isEmpty (A : Type) :
2017-08-07 16:49:46 +02:00
FSet A -> Bool.
Proof.
simple refine (FSet_rec _ _ _ true (fun _ => false) andb _ _ _ _ _)
; try eauto with bool_lattice_hints typeclass_instances.
intros ; symmetry ; eauto with lattice_hints typeclass_instances.
2017-08-07 22:13:42 +02:00
Defined.
2017-08-08 17:00:30 +02:00
Definition single_product {A B : Type} (a : A) : FSet B -> FSet (A * B).
2017-08-07 23:15:25 +02:00
Proof.
hrecursion.
- apply .
- intro b.
apply {|(a, b)|}.
2017-08-08 15:29:50 +02:00
- apply ().
2017-08-07 23:15:25 +02:00
- intros X Y Z ; apply assoc.
- intros X Y ; apply comm.
- intros ; apply nl.
- intros ; apply nr.
- intros ; apply idem.
Defined.
2017-08-08 17:00:30 +02:00
Definition product {A B : Type} : FSet A -> FSet B -> FSet (A * B).
2017-08-07 22:13:42 +02:00
Proof.
intros X Y.
hrecursion X.
- apply .
- intro a.
2017-08-07 23:15:25 +02:00
apply (single_product a Y).
2017-08-08 15:29:50 +02:00
- apply ().
2017-08-07 22:13:42 +02:00
- intros ; apply assoc.
- intros ; apply comm.
- intros ; apply nl.
- intros ; apply nr.
- intros ; apply union_idem.
2017-08-07 22:13:42 +02:00
Defined.
2017-08-08 17:00:30 +02:00
Global Instance fset_subset : forall A, hasSubset (FSet A).
2017-08-08 15:29:50 +02:00
Proof.
2017-08-08 17:00:30 +02:00
intros A X Y.
2017-08-08 15:29:50 +02:00
hrecursion X.
- exists Unit.
exact _.
- intros a.
apply (a Y).
- intros X1 X2.
exists (prod X1 X2).
exact _.
- intros.
apply path_trunctype ; apply equiv_prod_assoc.
- intros.
apply path_trunctype ; apply equiv_prod_symm.
- intros.
apply path_trunctype ; apply prod_unit_l.
- intros.
apply path_trunctype ; apply prod_unit_r.
- intros a'.
apply path_iff_hprop ; cbn.
* intros [p1 p2]. apply p1.
* intros p.
split ; apply p.
Defined.
2017-08-08 17:00:30 +02:00
End operations.