2017-05-24 13:54:00 +02:00
|
|
|
Require Import HoTT HitTactics.
|
|
|
|
Require Import definition operations.
|
2017-05-23 16:30:31 +02:00
|
|
|
|
2017-05-24 13:54:00 +02:00
|
|
|
Section properties.
|
2017-05-23 16:30:31 +02:00
|
|
|
|
2017-05-24 13:54:00 +02:00
|
|
|
Context {A : Type}.
|
|
|
|
Context {A_deceq : DecidablePaths A}.
|
2017-05-23 22:58:35 +02:00
|
|
|
|
2017-05-24 13:54:00 +02:00
|
|
|
(** union properties *)
|
|
|
|
Theorem union_idem : forall x: FSet A, U x x = x.
|
2017-05-23 16:30:31 +02:00
|
|
|
Proof.
|
2017-05-26 12:28:07 +02:00
|
|
|
hinduction;
|
2017-05-24 13:54:00 +02:00
|
|
|
try (intros ; apply set_path2) ; cbn.
|
|
|
|
- apply nl.
|
|
|
|
- apply idem.
|
|
|
|
- intros x y P Q.
|
|
|
|
rewrite assoc.
|
|
|
|
rewrite (comm x y).
|
|
|
|
rewrite <- (assoc y x x).
|
|
|
|
rewrite P.
|
|
|
|
rewrite (comm y x).
|
|
|
|
rewrite <- (assoc x y y).
|
|
|
|
rewrite Q.
|
2017-05-23 16:30:31 +02:00
|
|
|
reflexivity.
|
|
|
|
Defined.
|
|
|
|
|
2017-05-26 12:28:07 +02:00
|
|
|
|
2017-05-24 13:54:00 +02:00
|
|
|
(** isIn properties *)
|
2017-05-23 16:30:31 +02:00
|
|
|
Lemma isIn_singleton_eq (a b: A) : isIn a (L b) = true -> a = b.
|
2017-05-26 12:28:07 +02:00
|
|
|
Proof. unfold isIn. simpl.
|
2017-05-23 16:30:31 +02:00
|
|
|
destruct (dec (a = b)). intro. apply p.
|
2017-05-26 12:28:07 +02:00
|
|
|
intro X.
|
2017-05-23 16:30:31 +02:00
|
|
|
contradiction (false_ne_true X).
|
|
|
|
Defined.
|
|
|
|
|
|
|
|
Lemma isIn_empty_false (a: A) : isIn a E = true -> Empty.
|
2017-05-26 12:28:07 +02:00
|
|
|
Proof.
|
2017-05-23 16:30:31 +02:00
|
|
|
cbv. intro X.
|
2017-05-26 12:28:07 +02:00
|
|
|
contradiction (false_ne_true X).
|
2017-05-23 16:30:31 +02:00
|
|
|
Defined.
|
|
|
|
|
2017-05-26 12:28:07 +02:00
|
|
|
Lemma isIn_union (a: A) (X Y: FSet A) :
|
2017-05-23 16:30:31 +02:00
|
|
|
isIn a (U X Y) = (isIn a X || isIn a Y)%Bool.
|
2017-05-26 12:28:07 +02:00
|
|
|
Proof. reflexivity. Qed.
|
2017-05-23 16:30:31 +02:00
|
|
|
|
2017-05-24 13:54:00 +02:00
|
|
|
(** comprehension properties *)
|
|
|
|
Lemma comprehension_false Y : comprehension (fun a => isIn a E) 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.
|
2017-05-23 16:30:31 +02:00
|
|
|
|
|
|
|
Theorem comprehension_or : forall ϕ ψ (x: FSet A),
|
2017-05-26 12:28:07 +02:00
|
|
|
comprehension (fun a => orb (ϕ a) (ψ a)) x = U (comprehension ϕ x)
|
2017-05-23 16:30:31 +02:00
|
|
|
(comprehension ψ x).
|
|
|
|
Proof.
|
|
|
|
intros ϕ ψ.
|
2017-05-26 12:28:07 +02:00
|
|
|
hinduction; try (intros; apply set_path2).
|
2017-05-23 16:30:31 +02:00
|
|
|
- cbn. symmetry ; apply nl.
|
|
|
|
- cbn. intros.
|
|
|
|
destruct (ϕ a) ; destruct (ψ a) ; symmetry.
|
|
|
|
* apply idem.
|
2017-05-26 12:28:07 +02:00
|
|
|
* apply nr.
|
2017-05-23 16:30:31 +02:00
|
|
|
* apply nl.
|
|
|
|
* apply nl.
|
|
|
|
- simpl. intros x y P Q.
|
2017-05-26 12:28:07 +02:00
|
|
|
cbn.
|
2017-05-23 16:30:31 +02:00
|
|
|
rewrite P.
|
|
|
|
rewrite Q.
|
|
|
|
rewrite <- assoc.
|
|
|
|
rewrite (assoc (comprehension ψ x)).
|
|
|
|
rewrite (comm (comprehension ψ x) (comprehension ϕ y)).
|
|
|
|
rewrite <- assoc.
|
|
|
|
rewrite <- assoc.
|
|
|
|
reflexivity.
|
|
|
|
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-05-23 16:30:31 +02:00
|
|
|
- apply nl.
|
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 13:54:00 +02:00
|
|
|
(** intersection properties *)
|
2017-05-23 16:30:31 +02:00
|
|
|
Lemma intersection_0l: forall X: FSet A, intersection E X = E.
|
2017-05-26 12:28:07 +02:00
|
|
|
Proof.
|
|
|
|
hinduction;
|
2017-05-23 16:30:31 +02:00
|
|
|
try (intros ; apply set_path2).
|
|
|
|
- reflexivity.
|
|
|
|
- intro a.
|
|
|
|
reflexivity.
|
|
|
|
- unfold intersection.
|
|
|
|
intros x y P Q.
|
|
|
|
cbn.
|
|
|
|
rewrite P.
|
|
|
|
rewrite Q.
|
|
|
|
apply nl.
|
|
|
|
Defined.
|
|
|
|
|
2017-05-24 13:54:00 +02:00
|
|
|
Lemma intersection_0r (X: FSet A): intersection X E = E.
|
|
|
|
Proof. exact idpath. Defined.
|
2017-05-23 16:30:31 +02:00
|
|
|
|
2017-05-23 22:58:35 +02:00
|
|
|
Theorem intersection_La : forall (a : A) (X : FSet A),
|
|
|
|
intersection (L a) X = if isIn a X then L a else E.
|
|
|
|
Proof.
|
|
|
|
intro a.
|
2017-05-24 13:54:00 +02:00
|
|
|
hinduction; try (intros ; apply set_path2).
|
2017-05-23 22:58:35 +02:00
|
|
|
- reflexivity.
|
|
|
|
- intro b.
|
|
|
|
cbn.
|
2017-05-24 13:54:00 +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.
|
2017-05-23 22:58:35 +02:00
|
|
|
- unfold intersection.
|
|
|
|
intros x y P Q.
|
|
|
|
cbn.
|
|
|
|
rewrite P.
|
|
|
|
rewrite Q.
|
|
|
|
destruct (isIn a x) ; destruct (isIn a y).
|
|
|
|
* apply idem.
|
|
|
|
* apply nr.
|
2017-05-26 12:28:07 +02:00
|
|
|
* apply nl.
|
2017-05-23 22:58:35 +02:00
|
|
|
* apply nl.
|
|
|
|
Defined.
|
|
|
|
|
2017-05-24 13:54:00 +02:00
|
|
|
Lemma intersection_comm X Y: intersection X Y = intersection Y X.
|
|
|
|
Proof.
|
|
|
|
hrecursion X; try (intros; apply set_path2).
|
|
|
|
- cbn. unfold intersection. apply comprehension_false.
|
|
|
|
- cbn. unfold intersection. intros a.
|
|
|
|
hrecursion Y; try (intros; apply set_path2).
|
|
|
|
+ cbn. reflexivity.
|
|
|
|
+ cbn. intros b.
|
|
|
|
destruct (dec (a = b)) as [pa|npa].
|
|
|
|
* rewrite pa.
|
|
|
|
destruct (dec (b = b)) as [|nb]; [reflexivity|].
|
|
|
|
by contradiction nb.
|
|
|
|
* destruct (dec (b = a)) as [pb|]; [|reflexivity].
|
|
|
|
by contradiction npa.
|
|
|
|
+ cbn -[isIn]. intros Y1 Y2 IH1 IH2.
|
2017-05-26 12:28:07 +02:00
|
|
|
rewrite IH1.
|
2017-05-24 13:54:00 +02:00
|
|
|
rewrite IH2.
|
|
|
|
symmetry.
|
|
|
|
apply (comprehension_or (fun a => isIn a Y1) (fun a => isIn a Y2) (L a)).
|
|
|
|
- intros X1 X2 IH1 IH2.
|
|
|
|
cbn.
|
|
|
|
unfold intersection in *.
|
|
|
|
rewrite <- IH1.
|
2017-05-26 12:28:07 +02:00
|
|
|
rewrite <- IH2.
|
2017-05-24 13:54:00 +02:00
|
|
|
apply comprehension_or.
|
|
|
|
Defined.
|
|
|
|
|
|
|
|
Theorem intersection_idem : forall (X : FSet A), intersection X X = X.
|
|
|
|
Proof.
|
|
|
|
hinduction; try (intros; apply set_path2).
|
|
|
|
- reflexivity.
|
|
|
|
- intro a.
|
|
|
|
destruct (dec (a = a)).
|
|
|
|
* reflexivity.
|
|
|
|
* contradiction (n idpath).
|
|
|
|
- intros X Y IHX IHY.
|
|
|
|
unfold intersection in *.
|
|
|
|
rewrite comprehension_or.
|
|
|
|
rewrite comprehension_or.
|
|
|
|
rewrite IHX.
|
|
|
|
rewrite IHY.
|
|
|
|
rewrite comprehension_subset.
|
|
|
|
rewrite (comm X).
|
|
|
|
rewrite comprehension_subset.
|
|
|
|
reflexivity.
|
|
|
|
Defined.
|
|
|
|
|
|
|
|
(** assorted lattice laws *)
|
2017-05-26 12:28:07 +02:00
|
|
|
Lemma distributive_La (z : FSet A) (a : A) : forall Y : FSet A,
|
2017-05-24 18:28:24 +02:00
|
|
|
intersection (U (L a) z) Y = U (intersection (L a) Y) (intersection z Y).
|
2017-05-23 23:24:22 +02:00
|
|
|
Proof.
|
2017-05-24 13:54:00 +02:00
|
|
|
hinduction; try (intros ; apply set_path2) ; cbn.
|
2017-05-23 23:24:22 +02:00
|
|
|
- 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.
|
2017-05-24 13:54:00 +02:00
|
|
|
rewrite P. rewrite Q.
|
2017-05-23 23:24:22 +02:00
|
|
|
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.
|
|
|
|
|
2017-05-24 11:47:30 +02:00
|
|
|
Theorem distributive_intersection_U (X1 X2 Y : FSet A) :
|
2017-05-23 23:24:22 +02:00
|
|
|
intersection (U X1 X2) Y = U (intersection X1 Y) (intersection X2 Y).
|
|
|
|
Proof.
|
2017-05-24 13:54:00 +02:00
|
|
|
hinduction X1; try (intros ; apply set_path2) ; cbn.
|
2017-05-23 23:24:22 +02:00
|
|
|
- rewrite intersection_0l.
|
|
|
|
rewrite nl.
|
|
|
|
rewrite nl.
|
|
|
|
reflexivity.
|
|
|
|
- intro a.
|
|
|
|
rewrite intersection_La.
|
2017-05-24 12:10:39 +02:00
|
|
|
rewrite distributive_La.
|
2017-05-23 23:24:22 +02:00
|
|
|
rewrite intersection_La.
|
|
|
|
reflexivity.
|
|
|
|
- intros Z1 Z2 P Q.
|
|
|
|
unfold intersection in *.
|
|
|
|
cbn.
|
|
|
|
rewrite comprehension_or.
|
|
|
|
rewrite comprehension_or.
|
|
|
|
reflexivity.
|
|
|
|
Defined.
|
2017-05-24 13:54:00 +02:00
|
|
|
|
2017-05-23 22:58:35 +02:00
|
|
|
Theorem intersection_isIn : forall a (x y: FSet A),
|
|
|
|
isIn a (intersection x y) = andb (isIn a x) (isIn a y).
|
|
|
|
Proof.
|
|
|
|
intros a x y.
|
2017-05-24 13:54:00 +02:00
|
|
|
hinduction x; try (intros ; apply set_path2) ; cbn.
|
2017-05-23 22:58:35 +02:00
|
|
|
- rewrite intersection_0l.
|
|
|
|
reflexivity.
|
|
|
|
- intro b.
|
|
|
|
rewrite intersection_La.
|
|
|
|
destruct (dec (a = b)) ; cbn.
|
|
|
|
* rewrite p.
|
|
|
|
destruct (isIn b y).
|
|
|
|
+ cbn.
|
2017-05-24 13:54:00 +02:00
|
|
|
destruct (dec (b = b)); [reflexivity|].
|
|
|
|
by contradiction n.
|
2017-05-23 22:58:35 +02:00
|
|
|
+ reflexivity.
|
|
|
|
* destruct (isIn b y).
|
|
|
|
+ cbn.
|
2017-05-24 13:54:00 +02:00
|
|
|
destruct (dec (a = b)); [|reflexivity].
|
|
|
|
by contradiction n.
|
2017-05-23 22:58:35 +02:00
|
|
|
+ reflexivity.
|
|
|
|
- intros X1 X2 P Q.
|
2017-05-24 12:10:39 +02:00
|
|
|
rewrite distributive_intersection_U.
|
2017-05-23 22:58:35 +02:00
|
|
|
cbn.
|
|
|
|
rewrite P.
|
|
|
|
rewrite Q.
|
2017-05-26 12:28:07 +02:00
|
|
|
destruct (isIn a X1) ; destruct (isIn a X2) ; destruct (isIn a y) ;
|
2017-05-24 18:28:24 +02:00
|
|
|
reflexivity.
|
2017-05-23 22:58:35 +02:00
|
|
|
Defined.
|
2017-05-23 16:30:31 +02:00
|
|
|
|
|
|
|
|
|
|
|
|
2017-05-23 22:58:35 +02:00
|
|
|
Theorem intersection_assoc (X Y Z: FSet A) :
|
|
|
|
intersection X (intersection Y Z) = intersection (intersection X Y) Z.
|
|
|
|
Proof.
|
2017-05-24 13:54:00 +02:00
|
|
|
hinduction X; try (intros ; apply set_path2).
|
2017-05-23 22:58:35 +02:00
|
|
|
- cbn.
|
|
|
|
rewrite intersection_0l.
|
|
|
|
rewrite intersection_0l.
|
|
|
|
rewrite intersection_0l.
|
2017-05-23 16:30:31 +02:00
|
|
|
reflexivity.
|
2017-05-23 22:58:35 +02:00
|
|
|
- intros a.
|
2017-05-23 16:30:31 +02:00
|
|
|
cbn.
|
|
|
|
rewrite intersection_La.
|
|
|
|
rewrite intersection_La.
|
|
|
|
rewrite intersection_isIn.
|
2017-05-23 22:58:35 +02:00
|
|
|
destruct (isIn a Y).
|
2017-05-23 16:30:31 +02:00
|
|
|
* rewrite intersection_La.
|
2017-05-23 22:58:35 +02:00
|
|
|
destruct (isIn a Z).
|
2017-05-23 16:30:31 +02:00
|
|
|
+ reflexivity.
|
|
|
|
+ reflexivity.
|
2017-05-23 22:58:35 +02:00
|
|
|
* rewrite intersection_0l.
|
2017-05-26 12:28:07 +02:00
|
|
|
reflexivity.
|
2017-05-23 16:30:31 +02:00
|
|
|
- unfold intersection. cbn.
|
2017-05-23 22:58:35 +02:00
|
|
|
intros X1 X2 P Q.
|
2017-05-23 16:30:31 +02:00
|
|
|
rewrite comprehension_or.
|
|
|
|
rewrite P.
|
|
|
|
rewrite Q.
|
|
|
|
rewrite comprehension_or.
|
|
|
|
cbn.
|
|
|
|
rewrite comprehension_or.
|
|
|
|
reflexivity.
|
2017-05-23 22:58:35 +02:00
|
|
|
Defined.
|
2017-05-23 21:31:45 +02:00
|
|
|
|
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.
|
|
|
|
rewrite comprehension_or.
|
|
|
|
rewrite comprehension_or.
|
|
|
|
rewrite P.
|
|
|
|
rewrite Q.
|
|
|
|
rewrite comprehension_subset.
|
|
|
|
rewrite (comm X1).
|
|
|
|
rewrite comprehension_subset.
|
|
|
|
reflexivity.
|
|
|
|
Defined.
|
|
|
|
|
2017-05-26 12:28:07 +02:00
|
|
|
|
2017-05-24 11:47:30 +02:00
|
|
|
Theorem distributive_U_int (X1 X2 Y : FSet A) :
|
|
|
|
U (intersection X1 X2) Y = intersection (U X1 Y) (U X2 Y).
|
|
|
|
Proof.
|
2017-05-24 13:54:00 +02:00
|
|
|
hinduction X1; try (intros ; apply set_path2) ; cbn.
|
2017-05-24 11:47:30 +02:00
|
|
|
- rewrite intersection_0l.
|
|
|
|
rewrite nl.
|
2017-05-24 13:54:00 +02:00
|
|
|
unfold intersection.
|
2017-05-24 11:47:30 +02:00
|
|
|
rewrite comprehension_all.
|
|
|
|
pose (intersection_comm Y X2).
|
|
|
|
unfold intersection in p.
|
|
|
|
rewrite p.
|
|
|
|
rewrite comprehension_subset.
|
|
|
|
reflexivity.
|
2017-05-24 13:54:00 +02:00
|
|
|
- intros. unfold intersection. (* TODO isIn is simplified too much *)
|
2017-05-24 11:47:30 +02:00
|
|
|
rewrite comprehension_or.
|
|
|
|
rewrite comprehension_or.
|
2017-05-24 13:54:00 +02:00
|
|
|
(* rewrite intersection_La. *)
|
2017-05-24 11:47:30 +02:00
|
|
|
admit.
|
|
|
|
- unfold intersection.
|
|
|
|
cbn.
|
|
|
|
intros Z1 Z2 P Q.
|
|
|
|
rewrite comprehension_or.
|
2017-05-26 12:28:07 +02:00
|
|
|
assert (U (U (comprehension (fun a : A => isIn a Z1) X2)
|
2017-05-24 18:28:24 +02:00
|
|
|
(comprehension (fun a : A => isIn a Z2) X2))
|
2017-05-26 12:28:07 +02:00
|
|
|
Y = U (U (comprehension (fun a : A => isIn a Z1) X2)
|
2017-05-24 18:28:24 +02:00
|
|
|
(comprehension (fun a : A => isIn a Z2) X2))
|
2017-05-24 11:47:30 +02:00
|
|
|
(U Y Y)).
|
|
|
|
rewrite (union_idem Y).
|
|
|
|
reflexivity.
|
|
|
|
rewrite X.
|
|
|
|
rewrite <- assoc.
|
|
|
|
rewrite (assoc (comprehension (fun a : A => isIn a Z2) X2)).
|
|
|
|
rewrite Q.
|
2017-05-26 12:28:07 +02:00
|
|
|
rewrite
|
2017-05-24 18:28:24 +02:00
|
|
|
(comm (U (comprehension (fun a : A => (isIn a Z2 || isIn a Y)%Bool) X2)
|
|
|
|
(comprehension (fun a : A => (isIn a Z2 || isIn a Y)%Bool) Y)) Y).
|
2017-05-24 11:47:30 +02:00
|
|
|
rewrite assoc.
|
|
|
|
rewrite P.
|
|
|
|
rewrite <- assoc.
|
|
|
|
rewrite (assoc (comprehension (fun a : A => (isIn a Z1 || isIn a Y)%Bool) Y)).
|
|
|
|
rewrite (comm (comprehension (fun a : A => (isIn a Z1 || isIn a Y)%Bool) Y)).
|
|
|
|
rewrite <- assoc.
|
|
|
|
rewrite assoc.
|
|
|
|
enough (C : (U (comprehension (fun a : A => (isIn a Z1 || isIn a Y)%Bool) X2)
|
2017-05-26 12:28:07 +02:00
|
|
|
(comprehension (fun a : A => (isIn a Z2 || isIn a Y)%Bool) X2))
|
2017-05-24 18:28:24 +02:00
|
|
|
= (comprehension (fun a : A => (isIn a Z1 || isIn a Z2 || isIn a Y)%Bool) X2)).
|
2017-05-26 12:28:07 +02:00
|
|
|
rewrite C.
|
2017-05-24 11:47:30 +02:00
|
|
|
enough (D : (U (comprehension (fun a : A => (isIn a Z1 || isIn a Y)%Bool) Y)
|
2017-05-26 12:28:07 +02:00
|
|
|
(comprehension (fun a : A => (isIn a Z2 || isIn a Y)%Bool) Y))
|
2017-05-24 18:28:24 +02:00
|
|
|
= (comprehension (fun a : A => (isIn a Z1 || isIn a Z2 || isIn a Y)%Bool) Y)).
|
2017-05-24 11:47:30 +02:00
|
|
|
rewrite D.
|
|
|
|
reflexivity.
|
2017-05-24 12:10:39 +02:00
|
|
|
* repeat (rewrite comprehension_or).
|
2017-05-24 11:47:30 +02:00
|
|
|
rewrite <- assoc.
|
|
|
|
rewrite (comm (comprehension (fun a : A => isIn a Y) Y)).
|
|
|
|
rewrite <- (assoc (comprehension (fun a : A => isIn a Z2) Y)).
|
|
|
|
rewrite union_idem.
|
|
|
|
rewrite assoc.
|
|
|
|
reflexivity.
|
2017-05-24 13:54:00 +02:00
|
|
|
* repeat (rewrite comprehension_or).
|
2017-05-24 11:47:30 +02:00
|
|
|
rewrite <- assoc.
|
|
|
|
rewrite (comm (comprehension (fun a : A => isIn a Y) X2)).
|
|
|
|
rewrite <- (assoc (comprehension (fun a : A => isIn a Z2) X2)).
|
|
|
|
rewrite union_idem.
|
|
|
|
rewrite assoc.
|
|
|
|
reflexivity.
|
|
|
|
Admitted.
|
2017-05-23 21:31:45 +02:00
|
|
|
|
2017-05-24 12:10:39 +02:00
|
|
|
Theorem absorb_0 (X Y : FSet A) : U X (intersection X Y) = X.
|
|
|
|
Proof.
|
2017-05-24 13:54:00 +02:00
|
|
|
hinduction X; try (intros ; apply set_path2) ; cbn.
|
2017-05-24 12:10:39 +02:00
|
|
|
- rewrite nl.
|
|
|
|
apply intersection_0l.
|
|
|
|
- intro a.
|
|
|
|
rewrite intersection_La.
|
|
|
|
destruct (isIn a Y).
|
|
|
|
* apply union_idem.
|
|
|
|
* apply nr.
|
|
|
|
- intros X1 X2 P Q.
|
|
|
|
rewrite distributive_intersection_U.
|
|
|
|
rewrite <- assoc.
|
|
|
|
rewrite (comm X2).
|
|
|
|
rewrite assoc.
|
|
|
|
rewrite assoc.
|
|
|
|
rewrite P.
|
|
|
|
rewrite <- assoc.
|
|
|
|
rewrite (comm _ X2).
|
|
|
|
rewrite Q.
|
|
|
|
reflexivity.
|
|
|
|
Defined.
|
|
|
|
|
|
|
|
Theorem absorb_1 (X Y : FSet A) : intersection X (U X Y) = X.
|
|
|
|
Proof.
|
2017-05-24 13:54:00 +02:00
|
|
|
hrecursion X; try (intros ; apply set_path2).
|
2017-05-24 12:10:39 +02:00
|
|
|
- cbn.
|
|
|
|
rewrite nl.
|
|
|
|
apply comprehension_false.
|
|
|
|
- intro a.
|
|
|
|
rewrite intersection_La.
|
|
|
|
destruct (dec (a = a)).
|
|
|
|
* destruct (isIn a Y).
|
|
|
|
+ apply union_idem.
|
|
|
|
+ apply nr.
|
|
|
|
* contradiction (n idpath).
|
|
|
|
- intros X1 X2 P Q.
|
|
|
|
cbn in *.
|
|
|
|
symmetry.
|
|
|
|
rewrite <- P.
|
|
|
|
rewrite <- Q.
|
2017-05-24 13:54:00 +02:00
|
|
|
Admitted.
|
2017-05-24 18:28:24 +02:00
|
|
|
|
|
|
|
|
|
|
|
(* Properties about subset relation. *)
|
2017-05-26 12:28:07 +02:00
|
|
|
Lemma subsect_intersection `{Funext} (X Y : FSet A) :
|
|
|
|
X ⊆ Y = true -> U X Y = Y.
|
2017-05-24 18:28:24 +02:00
|
|
|
Proof.
|
|
|
|
hinduction X; try (intros; apply path_forall; intro; apply set_path2).
|
|
|
|
- intros. apply nl.
|
2017-05-26 12:28:07 +02:00
|
|
|
- intros a. hinduction Y;
|
2017-05-24 18:28:24 +02:00
|
|
|
try (intros; apply path_forall; intro; apply set_path2).
|
|
|
|
(*intros. apply equiv_hprop_allpath.*)
|
|
|
|
+ intro. cbn. contradiction (false_ne_true).
|
|
|
|
+ intros. destruct (dec (a = a0)).
|
|
|
|
rewrite p; apply idem.
|
|
|
|
contradiction (false_ne_true).
|
|
|
|
+ intros X1 X2 IH1 IH2.
|
|
|
|
intro Ho.
|
|
|
|
destruct (isIn a X1);
|
|
|
|
destruct (isIn a X2).
|
|
|
|
specialize (IH1 idpath).
|
|
|
|
specialize (IH2 idpath).
|
|
|
|
rewrite assoc. rewrite IH1. reflexivity.
|
|
|
|
specialize (IH1 idpath).
|
|
|
|
rewrite assoc. rewrite IH1. reflexivity.
|
|
|
|
specialize (IH2 idpath).
|
2017-05-26 12:28:07 +02:00
|
|
|
rewrite assoc. rewrite (comm (L a)). rewrite <- assoc. rewrite IH2.
|
2017-05-24 18:28:24 +02:00
|
|
|
reflexivity.
|
2017-05-26 12:28:07 +02:00
|
|
|
cbn in Ho. contradiction (false_ne_true).
|
|
|
|
- intros X1 X2 IH1 IH2 G.
|
2017-05-24 18:28:24 +02:00
|
|
|
destruct (subset X1 Y);
|
|
|
|
destruct (subset X2 Y).
|
|
|
|
specialize (IH1 idpath).
|
|
|
|
specialize (IH2 idpath).
|
|
|
|
rewrite <- assoc. rewrite IH2. rewrite IH1. reflexivity.
|
|
|
|
specialize (IH1 idpath).
|
|
|
|
apply IH2 in G.
|
|
|
|
rewrite <- assoc. rewrite G. rewrite IH1. reflexivity.
|
|
|
|
specialize (IH2 idpath).
|
|
|
|
apply IH1 in G.
|
|
|
|
rewrite <- assoc. rewrite IH2. rewrite G. reflexivity.
|
|
|
|
specialize (IH1 G). specialize (IH2 G).
|
|
|
|
rewrite <- assoc. rewrite IH2. rewrite IH1. reflexivity.
|
|
|
|
Defined.
|
|
|
|
|
2017-05-26 12:28:07 +02:00
|
|
|
Theorem
|
|
|
|
|
|
|
|
End properties.
|