2017-05-24 13:54:00 +02:00
|
|
|
Require Import HoTT HitTactics.
|
2017-06-19 21:32:55 +02:00
|
|
|
Require Export definition operations Ext Lattice.
|
2017-05-23 16:30:31 +02:00
|
|
|
|
2017-06-19 21:32:55 +02:00
|
|
|
(* Lemmas relating operations to the membership predicate *)
|
|
|
|
Section operations_isIn.
|
2017-06-20 15:08:52 +02:00
|
|
|
|
2017-06-19 21:32:55 +02:00
|
|
|
Context {A : Type} `{DecidablePaths A}.
|
2017-06-20 15:08:52 +02:00
|
|
|
|
2017-06-19 21:32:55 +02:00
|
|
|
Lemma ext `{Funext} : forall (S T : FSet A), (forall a, isIn a S = isIn a T) -> S = T.
|
|
|
|
Proof.
|
|
|
|
apply fset_ext.
|
|
|
|
Defined.
|
2017-05-23 16:30:31 +02:00
|
|
|
|
2017-06-19 21:32:55 +02:00
|
|
|
(* Union and membership *)
|
2017-06-20 15:08:52 +02:00
|
|
|
Lemma union_isIn (X Y : FSet A) (a : A) :
|
2017-06-19 21:32:55 +02:00
|
|
|
isIn a (U X Y) = orb (isIn a X) (isIn a Y).
|
|
|
|
Proof.
|
|
|
|
reflexivity.
|
2017-05-23 16:30:31 +02:00
|
|
|
Defined.
|
|
|
|
|
2017-06-19 21:32:55 +02:00
|
|
|
(* Intersection and membership. We need a bunch of supporting lemmas *)
|
|
|
|
Lemma intersection_0l: forall X: FSet A, intersection E X = E.
|
|
|
|
Proof.
|
|
|
|
hinduction;
|
|
|
|
try (intros ; apply set_path2).
|
|
|
|
- reflexivity.
|
|
|
|
- intro a.
|
|
|
|
reflexivity.
|
|
|
|
- intros x y P Q.
|
|
|
|
cbn.
|
|
|
|
rewrite P.
|
|
|
|
rewrite Q.
|
|
|
|
apply union_idem.
|
2017-05-23 16:30:31 +02:00
|
|
|
Defined.
|
|
|
|
|
2017-06-19 21:32:55 +02:00
|
|
|
Lemma intersection_0r (X : FSet A) : intersection X E = E.
|
2017-06-20 15:08:52 +02:00
|
|
|
Proof.
|
|
|
|
exact idpath.
|
|
|
|
Defined.
|
2017-05-23 16:30:31 +02:00
|
|
|
|
2017-06-19 21:32:55 +02:00
|
|
|
Lemma intersection_La (X : FSet A) (a : A) :
|
|
|
|
intersection (L a) X = if isIn a X then L a else E.
|
2017-05-24 13:54:00 +02:00
|
|
|
Proof.
|
2017-06-19 21:32:55 +02:00
|
|
|
hinduction X; try (intros ; apply set_path2).
|
|
|
|
- reflexivity.
|
|
|
|
- intro b.
|
2017-06-19 17:08:56 +02:00
|
|
|
cbn.
|
2017-06-19 21:32:55 +02:00
|
|
|
destruct (dec (a = b)) as [p|np].
|
|
|
|
* rewrite p.
|
|
|
|
destruct (dec (b = b)) as [|nb]; [reflexivity|].
|
|
|
|
by contradiction nb.
|
|
|
|
* destruct (dec (b = a)); [|reflexivity].
|
|
|
|
by contradiction np.
|
|
|
|
- unfold intersection.
|
|
|
|
intros X Y P Q.
|
|
|
|
cbn.
|
|
|
|
rewrite P.
|
|
|
|
rewrite Q.
|
|
|
|
destruct (isIn a X) ; destruct (isIn a Y).
|
|
|
|
* apply union_idem.
|
|
|
|
* apply nr.
|
|
|
|
* apply nl.
|
|
|
|
* apply union_idem.
|
2017-05-24 13:54:00 +02:00
|
|
|
Defined.
|
2017-05-23 16:30:31 +02:00
|
|
|
|
2017-06-19 21:32:55 +02:00
|
|
|
Lemma comprehension_or : forall ϕ ψ (x: FSet A),
|
2017-05-23 16:30:31 +02:00
|
|
|
comprehension (fun a => orb (ϕ a) (ψ a)) x = U (comprehension ϕ x)
|
|
|
|
(comprehension ψ x).
|
|
|
|
Proof.
|
|
|
|
intros ϕ ψ.
|
2017-05-24 13:54:00 +02:00
|
|
|
hinduction; try (intros; apply set_path2).
|
2017-06-19 21:06:17 +02:00
|
|
|
- cbn. apply (union_idem _)^.
|
2017-05-23 16:30:31 +02:00
|
|
|
- cbn. intros.
|
|
|
|
destruct (ϕ a) ; destruct (ψ a) ; symmetry.
|
2017-06-19 21:06:17 +02:00
|
|
|
* apply union_idem.
|
2017-05-23 16:30:31 +02:00
|
|
|
* apply nr.
|
|
|
|
* apply nl.
|
2017-06-19 21:06:17 +02:00
|
|
|
* apply union_idem.
|
2017-05-23 16:30:31 +02:00
|
|
|
- simpl. intros x y P Q.
|
|
|
|
rewrite P.
|
|
|
|
rewrite Q.
|
|
|
|
rewrite <- assoc.
|
|
|
|
rewrite (assoc (comprehension ψ x)).
|
|
|
|
rewrite (comm (comprehension ψ x) (comprehension ϕ y)).
|
|
|
|
rewrite <- assoc.
|
|
|
|
rewrite <- assoc.
|
|
|
|
reflexivity.
|
|
|
|
Defined.
|
|
|
|
|
2017-06-19 21:32:55 +02:00
|
|
|
Lemma distributive_La (z : FSet A) (a : A) : forall Y : FSet A,
|
|
|
|
intersection (U (L a) z) Y = U (intersection (L a) Y) (intersection z Y).
|
|
|
|
Proof.
|
|
|
|
hinduction; try (intros ; apply set_path2) ; cbn.
|
|
|
|
- symmetry ; apply nl.
|
|
|
|
- intros b.
|
|
|
|
destruct (dec (b = a)) ; cbn.
|
|
|
|
* destruct (isIn b z).
|
|
|
|
+ rewrite union_idem.
|
|
|
|
reflexivity.
|
|
|
|
+ rewrite nr.
|
|
|
|
reflexivity.
|
|
|
|
* rewrite nl ; reflexivity.
|
|
|
|
- intros X1 X2 P Q.
|
|
|
|
rewrite P. rewrite Q.
|
|
|
|
rewrite <- assoc.
|
|
|
|
rewrite (comm (comprehension (fun a0 : A => isIn a0 z) X1)).
|
|
|
|
rewrite <- assoc.
|
|
|
|
rewrite assoc.
|
|
|
|
rewrite (comm (comprehension (fun a0 : A => isIn a0 z) X2)).
|
|
|
|
reflexivity.
|
|
|
|
Defined.
|
|
|
|
|
|
|
|
Lemma distributive_intersection_U (X1 X2 Y : FSet A) :
|
|
|
|
intersection (U X1 X2) Y = U (intersection X1 Y) (intersection X2 Y).
|
|
|
|
Proof.
|
|
|
|
hinduction X1; try (intros ; apply set_path2) ; cbn.
|
|
|
|
- rewrite intersection_0l.
|
|
|
|
rewrite nl.
|
|
|
|
rewrite nl.
|
|
|
|
reflexivity.
|
|
|
|
- intro a.
|
|
|
|
rewrite intersection_La.
|
|
|
|
rewrite distributive_La.
|
|
|
|
rewrite intersection_La.
|
|
|
|
reflexivity.
|
|
|
|
- intros Z1 Z2 P Q.
|
|
|
|
unfold intersection in *. simpl in *.
|
|
|
|
apply comprehension_or.
|
|
|
|
Defined.
|
|
|
|
|
|
|
|
Theorem intersection_isIn (X Y: FSet A) (a : A) :
|
|
|
|
isIn a (intersection X Y) = andb (isIn a X) (isIn a Y).
|
|
|
|
Proof.
|
|
|
|
hinduction X; try (intros ; apply set_path2) ; cbn.
|
|
|
|
- rewrite intersection_0l.
|
|
|
|
reflexivity.
|
|
|
|
- intro b.
|
|
|
|
rewrite intersection_La.
|
|
|
|
destruct (dec (a = b)) ; cbn.
|
|
|
|
* rewrite p.
|
|
|
|
destruct (isIn b Y).
|
|
|
|
+ cbn.
|
|
|
|
destruct (dec (b = b)); [reflexivity|].
|
|
|
|
by contradiction n.
|
|
|
|
+ reflexivity.
|
|
|
|
* destruct (isIn b Y).
|
|
|
|
+ cbn.
|
|
|
|
destruct (dec (a = b)); [|reflexivity].
|
|
|
|
by contradiction n.
|
|
|
|
+ reflexivity.
|
|
|
|
- intros X1 X2 P Q.
|
|
|
|
rewrite distributive_intersection_U. simpl.
|
|
|
|
rewrite P.
|
|
|
|
rewrite Q.
|
|
|
|
destruct (isIn a X1) ; destruct (isIn a X2) ; destruct (isIn a Y) ;
|
|
|
|
reflexivity.
|
|
|
|
Defined.
|
|
|
|
End operations_isIn.
|
|
|
|
|
|
|
|
(* Some suporting tactics *)
|
|
|
|
Ltac simplify_isIn :=
|
|
|
|
repeat (rewrite ?intersection_isIn ;
|
|
|
|
rewrite ?union_isIn).
|
|
|
|
|
|
|
|
Ltac toBool := try (intro) ;
|
|
|
|
intros ; apply ext ; intros ; simplify_isIn ; eauto with bool_lattice_hints.
|
|
|
|
|
|
|
|
Section SetLattice.
|
|
|
|
|
|
|
|
Context {A : Type}.
|
|
|
|
Context {A_deceq : DecidablePaths A}.
|
|
|
|
Context `{Funext}.
|
|
|
|
|
|
|
|
Instance fset_union_com : Commutative (@U A).
|
|
|
|
Proof.
|
|
|
|
toBool.
|
|
|
|
Defined.
|
|
|
|
|
|
|
|
Instance fset_intersection_com : Commutative intersection.
|
|
|
|
Proof.
|
|
|
|
toBool.
|
|
|
|
Defined.
|
|
|
|
|
|
|
|
Instance fset_union_assoc : Associative (@U A).
|
|
|
|
Proof.
|
|
|
|
toBool.
|
|
|
|
Defined.
|
|
|
|
|
|
|
|
Instance fset_intersection_assoc : Associative intersection.
|
|
|
|
Proof.
|
|
|
|
toBool.
|
|
|
|
Defined.
|
|
|
|
|
|
|
|
Instance fset_union_idem : Idempotent (@U A).
|
|
|
|
Proof. exact union_idem. Defined.
|
|
|
|
|
|
|
|
Instance fset_intersection_idem : Idempotent intersection.
|
|
|
|
Proof.
|
|
|
|
toBool.
|
|
|
|
Defined.
|
|
|
|
|
|
|
|
Instance fset_union_nl : NeutralL (@U A) (@E A).
|
|
|
|
Proof.
|
|
|
|
toBool.
|
|
|
|
Defined.
|
|
|
|
|
|
|
|
Instance fset_union_nr : NeutralR (@U A) (@E A).
|
|
|
|
Proof.
|
|
|
|
toBool.
|
|
|
|
Defined.
|
|
|
|
|
|
|
|
Instance fset_absorption_intersection_union : Absorption (@U A) intersection.
|
|
|
|
Proof.
|
|
|
|
toBool.
|
|
|
|
Defined.
|
|
|
|
|
|
|
|
Instance fset_absorption_union_intersection : Absorption intersection (@U A).
|
|
|
|
Proof.
|
|
|
|
toBool.
|
|
|
|
Defined.
|
|
|
|
|
|
|
|
Instance lattice_fset : Lattice (@U A) intersection (@E A) :=
|
|
|
|
{
|
|
|
|
commutative_min := _ ;
|
|
|
|
commutative_max := _ ;
|
|
|
|
associative_min := _ ;
|
|
|
|
associative_max := _ ;
|
|
|
|
idempotent_min := _ ;
|
|
|
|
idempotent_max := _ ;
|
|
|
|
neutralL_min := _ ;
|
|
|
|
neutralR_min := _ ;
|
|
|
|
absorption_min_max := _ ;
|
|
|
|
absorption_max_min := _
|
|
|
|
}.
|
|
|
|
|
|
|
|
End SetLattice.
|
|
|
|
|
|
|
|
(* Other properties *)
|
|
|
|
Section properties.
|
|
|
|
|
|
|
|
Context {A : Type}.
|
|
|
|
Context {A_deceq : DecidablePaths A}.
|
|
|
|
|
|
|
|
(** isIn properties *)
|
2017-06-20 15:08:52 +02:00
|
|
|
Lemma singleton_isIn (a b: A) : isIn a (L b) = true -> a = b.
|
|
|
|
Proof.
|
|
|
|
simpl.
|
|
|
|
destruct (dec (a = b)).
|
|
|
|
- intro.
|
|
|
|
apply p.
|
|
|
|
- intro X.
|
|
|
|
contradiction (false_ne_true X).
|
2017-06-19 21:32:55 +02:00
|
|
|
Defined.
|
|
|
|
|
2017-06-20 15:08:52 +02:00
|
|
|
Lemma empty_isIn (a: A) : isIn a E = false.
|
2017-06-19 21:32:55 +02:00
|
|
|
Proof.
|
2017-06-20 15:08:52 +02:00
|
|
|
reflexivity.
|
2017-06-19 21:32:55 +02:00
|
|
|
Defined.
|
|
|
|
|
|
|
|
(** comprehension properties *)
|
|
|
|
Lemma comprehension_false Y : comprehension (fun (_ : A) => false) Y = E.
|
|
|
|
Proof.
|
|
|
|
hrecursion Y; try (intros; apply set_path2).
|
|
|
|
- reflexivity.
|
|
|
|
- reflexivity.
|
|
|
|
- intros x y IHa IHb.
|
|
|
|
rewrite IHa.
|
|
|
|
rewrite IHb.
|
|
|
|
apply union_idem.
|
|
|
|
Defined.
|
|
|
|
|
2017-05-24 13:54:00 +02:00
|
|
|
Theorem comprehension_subset : forall ϕ (X : FSet A),
|
|
|
|
U (comprehension ϕ X) X = X.
|
2017-05-23 16:30:31 +02:00
|
|
|
Proof.
|
2017-05-24 13:54:00 +02:00
|
|
|
intros ϕ.
|
|
|
|
hrecursion; try (intros ; apply set_path2) ; cbn.
|
2017-06-19 21:06:17 +02:00
|
|
|
- apply union_idem.
|
2017-05-24 13:54:00 +02:00
|
|
|
- intro a.
|
|
|
|
destruct (ϕ a).
|
|
|
|
* apply union_idem.
|
|
|
|
* apply nl.
|
|
|
|
- intros X Y P Q.
|
|
|
|
rewrite assoc.
|
|
|
|
rewrite <- (assoc (comprehension ϕ X) (comprehension ϕ Y) X).
|
|
|
|
rewrite (comm (comprehension ϕ Y) X).
|
2017-05-23 16:30:31 +02:00
|
|
|
rewrite assoc.
|
|
|
|
rewrite P.
|
2017-05-24 13:54:00 +02:00
|
|
|
rewrite <- assoc.
|
2017-05-23 16:30:31 +02:00
|
|
|
rewrite Q.
|
|
|
|
reflexivity.
|
|
|
|
Defined.
|
|
|
|
|
2017-05-24 11:47:30 +02:00
|
|
|
Theorem comprehension_all : forall (X : FSet A),
|
|
|
|
comprehension (fun a => isIn a X) X = X.
|
|
|
|
Proof.
|
2017-05-24 13:54:00 +02:00
|
|
|
hinduction; try (intros ; apply set_path2).
|
2017-05-24 11:47:30 +02:00
|
|
|
- reflexivity.
|
|
|
|
- intro a.
|
|
|
|
destruct (dec (a = a)).
|
|
|
|
* reflexivity.
|
|
|
|
* contradiction (n idpath).
|
|
|
|
- intros X1 X2 P Q.
|
2017-05-25 15:11:41 +02:00
|
|
|
f_ap; (etransitivity; [ apply comprehension_or |]).
|
|
|
|
rewrite P. rewrite (comm X1).
|
|
|
|
apply comprehension_subset.
|
|
|
|
|
2017-05-24 11:47:30 +02:00
|
|
|
rewrite Q.
|
2017-05-25 15:11:41 +02:00
|
|
|
apply comprehension_subset.
|
2017-05-24 11:47:30 +02:00
|
|
|
Defined.
|
2017-05-23 16:30:31 +02:00
|
|
|
|
2017-06-19 21:32:55 +02:00
|
|
|
Theorem distributive_U_int `{Funext} (X1 X2 Y : FSet A) :
|
2017-05-24 11:47:30 +02:00
|
|
|
U (intersection X1 X2) Y = intersection (U X1 Y) (U X2 Y).
|
|
|
|
Proof.
|
2017-06-19 21:32:55 +02:00
|
|
|
toBool.
|
|
|
|
destruct (a ∈ X1), (a ∈ X2), (a ∈ Y); eauto.
|
|
|
|
Defined.
|
2017-05-23 21:31:45 +02:00
|
|
|
|
2017-06-19 21:06:17 +02:00
|
|
|
End properties.
|