HITs-Examples/FinSets.v

475 lines
13 KiB
Coq
Raw Normal View History

2017-01-02 13:08:36 +01:00
Require Export HoTT.
Require Import HitTactics.
2017-01-02 13:08:36 +01:00
Module Export FinSet.
Section FSet.
Variable A : Type.
2017-01-02 13:08:36 +01:00
Private Inductive FSet : Type :=
| E : FSet
| L : A -> FSet
| U : FSet -> FSet -> FSet.
2017-01-02 13:08:36 +01:00
Notation "{| x |}" := (L x).
Infix "" := U (at level 8, right associativity).
Notation "" := E.
2017-01-02 13:08:36 +01:00
Axiom assoc : forall (x y z : FSet ),
x (y z) = (x y) z.
2017-01-02 13:08:36 +01:00
Axiom comm : forall (x y : FSet),
x y = y x.
2017-01-02 13:08:36 +01:00
Axiom nl : forall (x : FSet),
x = x.
2017-01-02 13:08:36 +01:00
Axiom nr : forall (x : FSet),
x = x.
2017-01-02 13:08:36 +01:00
Axiom idem : forall (x : A),
{| x |} {|x|} = {|x|}.
Axiom trunc : IsHSet FSet.
Fixpoint FSet_rec
2017-01-02 13:08:36 +01:00
(P : Type)
(H: IsHSet P)
2017-01-02 13:08:36 +01:00
(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)
(x : FSet)
2017-01-02 13:08:36 +01:00
{struct x}
: P
:= (match x return _ -> _ -> _ -> _ -> _ -> _ -> P with
| E => fun _ => fun _ => fun _ => fun _ => fun _ => fun _ => e
| L a => fun _ => fun _ => fun _ => fun _ => fun _ => fun _ => l a
| U y z => fun _ => fun _ => fun _ => fun _ => fun _ => fun _ => u
(FSet_rec P H e l u assocP commP nlP nrP idemP y)
(FSet_rec P H e l u assocP commP nlP nrP idemP z)
end) assocP commP nlP nrP idemP H.
Axiom FSet_rec_beta_assoc : forall
2017-01-02 13:08:36 +01:00
(P : Type)
(H: IsHSet P)
2017-01-02 13:08:36 +01:00
(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)
(x y z : FSet),
ap (FSet_rec P H e l u assocP commP nlP nrP idemP) (assoc x y z)
2017-01-02 13:08:36 +01:00
=
(assocP (FSet_rec P H e l u assocP commP nlP nrP idemP x)
(FSet_rec P H e l u assocP commP nlP nrP idemP y)
(FSet_rec P H e l u assocP commP nlP nrP idemP z)
2017-01-02 13:08:36 +01:00
).
Axiom FSet_rec_beta_comm : forall
2017-01-02 13:08:36 +01:00
(P : Type)
(H: IsHSet P)
2017-01-02 13:08:36 +01:00
(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)
(x y : FSet),
ap (FSet_rec P H e l u assocP commP nlP nrP idemP) (comm x y)
2017-01-02 13:08:36 +01:00
=
(commP (FSet_rec P H e l u assocP commP nlP nrP idemP x)
(FSet_rec P H e l u assocP commP nlP nrP idemP y)
2017-01-02 13:08:36 +01:00
).
Axiom FSet_rec_beta_nl : forall
2017-01-02 13:08:36 +01:00
(P : Type)
(H: IsHSet P)
2017-01-02 13:08:36 +01:00
(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)
(x : FSet),
ap (FSet_rec P H e l u assocP commP nlP nrP idemP) (nl x)
2017-01-02 13:08:36 +01:00
=
(nlP (FSet_rec P H e l u assocP commP nlP nrP idemP x)
2017-01-02 13:08:36 +01:00
).
Axiom FSet_rec_beta_nr : forall
2017-01-02 13:08:36 +01:00
(P : Type)
(H: IsHSet P)
2017-01-02 13:08:36 +01:00
(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)
(x : FSet),
ap (FSet_rec P H e l u assocP commP nlP nrP idemP) (nr x)
2017-01-02 13:08:36 +01:00
=
(nrP (FSet_rec P H e l u assocP commP nlP nrP idemP x)
2017-01-02 13:08:36 +01:00
).
Axiom FSet_rec_beta_idem : forall
2017-01-02 13:08:36 +01:00
(P : Type)
(H: IsHSet P)
2017-01-02 13:08:36 +01:00
(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)
(x : A),
ap (FSet_rec P H e l u assocP commP nlP nrP idemP) (idem x)
2017-01-02 13:08:36 +01:00
=
idemP x.
(* Induction principle *)
Fixpoint FSet_ind
(P : FSet -> Type)
(H : forall a : FSet, IsHSet (P a))
(eP : P E)
(lP : forall a: A, P (L a))
(uP : forall (x y: FSet), P x -> P y -> P (U x y))
(assocP : forall (x y z : FSet)
(px: P x) (py: P y) (pz: P z),
assoc x y z #
(uP x (U y z) px (uP y z py pz))
=
(uP (U x y) z (uP x y px py) pz))
(commP : forall (x y: FSet) (px: P x) (py: P y),
comm x y #
uP x y px py = uP y x py px)
(nlP : forall (x : FSet) (px: P x),
nl x # uP E x eP px = px)
(nrP : forall (x : FSet) (px: P x),
nr x # uP x E px eP = px)
(idemP : forall (x : A),
idem x # uP (L x) (L x) (lP x) (lP x) = lP x)
(x : FSet)
{struct x}
: P x
:= (match x return _ -> _ -> _ -> _ -> _ -> _ -> P x with
| E => fun _ => fun _ => fun _ => fun _ => fun _ => fun _ => eP
| L a => fun _ => fun _ => fun _ => fun _ => fun _ => fun _ => lP a
| U y z => fun _ => fun _ => fun _ => fun _ => fun _ => fun _ => uP y z
(FSet_ind P H eP lP uP assocP commP nlP nrP idemP y)
(FSet_ind P H eP lP uP assocP commP nlP nrP idemP z)
end) H assocP commP nlP nrP idemP.
Axiom FSet_ind_beta_assoc : forall
(P : FSet -> Type)
(H : forall a : FSet, IsHSet (P a))
(eP : P E)
(lP : forall a: A, P (L a))
(uP : forall (x y: FSet), P x -> P y -> P (U x y))
(assocP : forall (x y z : FSet)
(px: P x) (py: P y) (pz: P z),
assoc x y z #
(uP x (U y z) px (uP y z py pz))
=
(uP (U x y) z (uP x y px py) pz))
(commP : forall (x y: FSet) (px: P x) (py: P y),
comm x y #
uP x y px py = uP y x py px)
(nlP : forall (x : FSet) (px: P x),
nl x # uP E x eP px = px)
(nrP : forall (x : FSet) (px: P x),
nr x # uP x E px eP = px)
(idemP : forall (x : A),
idem x # uP (L x) (L x) (lP x) (lP x) = lP x)
(x y z : FSet),
apD (FSet_ind P H eP lP uP assocP commP nlP nrP idemP)
(assoc x y z)
=
(assocP x y z
(FSet_ind P H eP lP uP assocP commP nlP nrP idemP x)
(FSet_ind P H eP lP uP assocP commP nlP nrP idemP y)
(FSet_ind P H eP lP uP assocP commP nlP nrP idemP z)
).
Axiom FSet_ind_beta_comm : forall
(P : FSet -> Type)
(H : forall a : FSet, IsHSet (P a))
(eP : P E)
(lP : forall a: A, P (L a))
(uP : forall (x y: FSet), P x -> P y -> P (U x y))
(assocP : forall (x y z : FSet)
(px: P x) (py: P y) (pz: P z),
assoc x y z #
(uP x (U y z) px (uP y z py pz))
=
(uP (U x y) z (uP x y px py) pz))
(commP : forall (x y : FSet) (px: P x) (py: P y),
comm x y #
uP x y px py = uP y x py px)
(nlP : forall (x : FSet) (px: P x),
nl x # uP E x eP px = px)
(nrP : forall (x : FSet) (px: P x),
nr x # uP x E px eP = px)
(idemP : forall (x : A),
idem x # uP (L x) (L x) (lP x) (lP x) = lP x)
(x y : FSet),
apD (FSet_ind P H eP lP uP assocP commP nlP nrP idemP) (comm x y)
=
(commP x y
(FSet_ind P H eP lP uP assocP commP nlP nrP idemP x)
(FSet_ind P H eP lP uP assocP commP nlP nrP idemP y)
).
Axiom FSet_ind_beta_nl : forall
(P : FSet -> Type)
(H : forall a : FSet, IsHSet (P a))
(eP : P E)
(lP : forall a: A, P (L a))
(uP : forall (x y: FSet), P x -> P y -> P (U x y))
(assocP : forall (x y z : FSet)
(px: P x) (py: P y) (pz: P z),
assoc x y z #
(uP x (U y z) px (uP y z py pz))
=
(uP (U x y) z (uP x y px py) pz))
(commP : forall (x y : FSet) (px: P x) (py: P y),
comm x y #
uP x y px py = uP y x py px)
(nlP : forall (x : FSet) (px: P x),
nl x # uP E x eP px = px)
(nrP : forall (x : FSet) (px: P x),
nr x # uP x E px eP = px)
(idemP : forall (x : A),
idem x # uP (L x) (L x) (lP x) (lP x) = lP x)
(x : FSet),
apD (FSet_ind P H eP lP uP assocP commP nlP nrP idemP) (nl x)
=
(nlP x (FSet_ind P H eP lP uP assocP commP nlP nrP idemP x)
).
Axiom FSet_ind_beta_nr : forall
(P : FSet -> Type)
(H : forall a : FSet, IsHSet (P a))
(eP : P E)
(lP : forall a: A, P (L a))
(uP : forall (x y: FSet), P x -> P y -> P (U x y))
(assocP : forall (x y z : FSet)
(px: P x) (py: P y) (pz: P z),
assoc x y z #
(uP x (U y z) px (uP y z py pz))
=
(uP (U x y) z (uP x y px py) pz))
(commP : forall (x y : FSet) (px: P x) (py: P y),
comm x y #
uP x y px py = uP y x py px)
(nlP : forall (x : FSet) (px: P x),
nl x # uP E x eP px = px)
(nrP : forall (x : FSet) (px: P x),
nr x # uP x E px eP = px)
(idemP : forall (x : A),
idem x # uP (L x) (L x) (lP x) (lP x) = lP x)
(x : FSet),
apD (FSet_ind P H eP lP uP assocP commP nlP nrP idemP) (nr x)
=
(nrP x (FSet_ind P H eP lP uP assocP commP nlP nrP idemP x)
).
Axiom FSet_ind_beta_idem : forall
(P : FSet -> Type)
(H : forall a : FSet, IsHSet (P a))
(eP : P E)
(lP : forall a: A, P (L a))
(uP : forall (x y: FSet), P x -> P y -> P (U x y))
(assocP : forall (x y z : FSet)
(px: P x) (py: P y) (pz: P z),
assoc x y z #
(uP x (U y z) px (uP y z py pz))
=
(uP (U x y) z (uP x y px py) pz))
(commP : forall (x y : FSet) (px: P x) (py: P y),
comm x y #
uP x y px py = uP y x py px)
(nlP : forall (x : FSet) (px: P x),
nl x # uP E x eP px = px)
(nrP : forall (x : FSet) (px: P x),
nr x # uP x E px eP = px)
(idemP : forall (x : A),
idem x # uP (L x) (L x) (lP x) (lP x) = lP x)
(x : A),
apD (FSet_ind P H eP lP uP assocP commP nlP nrP idemP) (idem x)
=
idemP x.
End FSet.
Instance FSet_recursion A : HitRecursion (FSet A) := {
indTy := _; recTy := _;
H_inductor := FSet_ind A; H_recursor := FSet_rec A }.
Arguments E {_}.
Arguments U {_} _ _.
Arguments L {_} _.
2017-01-02 13:08:36 +01:00
End FinSet.
Section functions.
Parameter A : Type.
2017-05-23 11:58:09 +02:00
Context (A_eqdec: forall (x y : A), Decidable (x = y)).
Definition deceq (x y : A) :=
if dec (x = y) then true else false.
2017-05-23 11:58:09 +02:00
Notation "{| x |}" := (L x).
Infix "" := U (at level 8, right associativity).
Notation "" := E.
(** Properties of union *)
Lemma union_idem : forall (X : FSet A), U X X = X.
Proof.
hinduction; try (intros; apply set_path2).
- apply nr.
- intros. apply idem.
- intros X Y HX HY. etransitivity.
rewrite assoc. rewrite (comm _ X Y). rewrite <- (assoc _ Y X X).
rewrite comm.
rewrite assoc. rewrite HX. rewrite HY. reflexivity.
rewrite comm. reflexivity.
Defined.
2017-05-23 11:58:09 +02:00
(** Membership predicate *)
Definition isIn : A -> FSet A -> Bool.
2017-01-02 13:08:36 +01:00
Proof.
intros a.
hrecursion.
- exact false.
- intro a'. apply (deceq a a').
- apply orb.
- intros x y z. compute. destruct x; reflexivity.
- intros x y. compute. destruct x, y; reflexivity.
- intros x. compute. reflexivity.
- intros x. compute. destruct x; reflexivity.
- intros a'. compute. destruct (A_eqdec a a'); reflexivity.
Defined.
2017-05-23 11:58:09 +02:00
Infix "" := isIn (at level 9, right associativity).
Lemma isIn_singleton_eq a b : a {| b |} = true -> a = b.
Proof. cbv.
destruct (A_eqdec a b). intro. apply p.
intro X.
2017-05-23 11:58:09 +02:00
contradiction (false_ne_true X).
Defined.
2017-05-23 11:58:09 +02:00
Lemma isIn_empty_false a : a = true -> Empty.
Proof.
cbv. intro X.
2017-05-23 11:58:09 +02:00
contradiction (false_ne_true X).
2017-01-02 13:08:36 +01:00
Defined.
2017-05-23 11:58:09 +02:00
Lemma isIn_union a X Y : a (X Y) = (a X || a Y)%Bool.
Proof. reflexivity. Qed.
2017-05-23 11:58:09 +02:00
(** Set comprehension *)
Definition comprehension :
(A -> Bool) -> FSet A -> FSet A.
2017-01-02 13:08:36 +01:00
Proof.
intros P.
hrecursion.
- apply E.
- intro a.
refine (if (P a) then L a else E).
- apply U.
- intros. cbv. apply assoc.
- intros. cbv. apply comm.
- intros. cbv. apply nl.
- intros. cbv. apply nr.
- intros. cbv.
destruct (P x); simpl.
+ apply idem.
+ apply nl.
2017-01-02 13:08:36 +01:00
Defined.
2017-05-23 11:58:09 +02:00
Lemma comprehension_false Y : comprehension (fun a => a ) Y = E.
Proof.
hrecursion Y; try (intros; apply set_path2).
- cbn. reflexivity.
- cbn. reflexivity.
- intros x y IHa IHb.
cbn.
rewrite IHa.
rewrite IHb.
rewrite nl.
reflexivity.
Defined.
Lemma comprehension_idem' `{Funext}:
2017-05-23 11:58:09 +02:00
forall (X:FSet A), forall Y, comprehension (fun x => x (U X Y)) X = X.
Proof.
hinduction.
all: try (intros; apply path_forall; intro; apply set_path2).
- intro Y. cbv. reflexivity.
- intros a Y. cbn.
unfold deceq;
destruct (dec (a = a)); simpl.
+ reflexivity.
+ contradiction n. reflexivity.
- intros X1 X2 IH1 IH2 Y.
cbn -[isIn].
f_ap.
+ rewrite <- assoc. apply (IH1 (U X2 Y)).
+ rewrite (comm _ X1 X2).
rewrite <- (assoc _ X2 X1 Y).
apply (IH2 (U X1 Y)).
Defined.
Lemma comprehension_idem `{Funext}:
2017-05-23 11:58:09 +02:00
forall (X:FSet A), comprehension (fun x => x X) X = X.
Proof.
intros X.
enough (comprehension (fun x : A => isIn x (U X X)) X = X).
rewrite (union_idem) in X0. assumption.
apply comprehension_idem'.
Defined.
2017-05-23 11:58:09 +02:00
(** Set intersection *)
Definition intersection :
FSet A -> FSet A -> FSet A.
2017-01-02 13:08:36 +01:00
Proof.
intros X Y.
apply (comprehension (fun (a : A) => isIn a X) Y).
2017-01-02 13:08:36 +01:00
Defined.
2017-05-23 11:58:09 +02:00
(** Subset ordering *)
Definition subset : forall (x : FSet A) (y : FSet A), Bool.
Proof. intros x y.
hrecursion x.
- apply true.
- intro a. apply (isIn a y).
2017-05-23 11:58:09 +02:00
- intros a b. apply (andb a b).
- intros a b c. compute. destruct a; reflexivity.
- intros a b. compute. destruct a, b; reflexivity.
- intros x. compute. reflexivity.
- intros x. compute. destruct x;reflexivity.
- intros a. simpl.
destruct (isIn a y); reflexivity.
Defined.
2017-05-23 11:58:09 +02:00
Infix "" := subset (at level 8, right associativity).
End functions.