mirror of
https://github.com/nmvdw/HITs-Examples
synced 2025-11-03 15:13:51 +01:00
Split the development into different directories
This commit is contained in:
111
FiniteSets/fsets/extensionality.v
Normal file
111
FiniteSets/fsets/extensionality.v
Normal file
@@ -0,0 +1,111 @@
|
||||
(** Extensionality of the FSets *)
|
||||
Require Import HoTT HitTactics.
|
||||
From representations Require Import definition.
|
||||
From fsets Require Import operations properties.
|
||||
|
||||
Section ext.
|
||||
Context {A : Type}.
|
||||
Context `{Univalence}.
|
||||
|
||||
Lemma subset_union_equiv
|
||||
: forall X Y : FSet A, subset X Y <~> U X Y = Y.
|
||||
Proof.
|
||||
intros X Y.
|
||||
eapply equiv_iff_hprop_uncurried.
|
||||
split.
|
||||
- apply subset_union.
|
||||
- intro HXY.
|
||||
rewrite <- HXY.
|
||||
apply subset_union_l.
|
||||
Defined.
|
||||
|
||||
Lemma subset_isIn (X Y : FSet A) :
|
||||
(forall (a : A), isIn a X -> isIn a Y)
|
||||
<~> (subset X Y).
|
||||
Proof.
|
||||
eapply equiv_iff_hprop_uncurried.
|
||||
split.
|
||||
- hinduction X ;
|
||||
try (intros; repeat (apply path_forall; intro); apply equiv_hprop_allpath ; apply _).
|
||||
+ intros ; reflexivity.
|
||||
+ intros a f.
|
||||
apply f.
|
||||
apply tr ; reflexivity.
|
||||
+ intros X1 X2 H1 H2 f.
|
||||
enough (subset X1 Y).
|
||||
enough (subset X2 Y).
|
||||
{ split. apply X. apply X0. }
|
||||
* apply H2.
|
||||
intros a Ha.
|
||||
apply f.
|
||||
apply tr.
|
||||
right.
|
||||
apply Ha.
|
||||
* apply H1.
|
||||
intros a Ha.
|
||||
apply f.
|
||||
apply tr.
|
||||
left.
|
||||
apply Ha.
|
||||
- hinduction X ;
|
||||
try (intros; repeat (apply path_forall; intro); apply equiv_hprop_allpath ; apply _).
|
||||
+ intros. contradiction.
|
||||
+ intros b f a.
|
||||
simple refine (Trunc_ind _ _) ; cbn.
|
||||
intro p.
|
||||
rewrite p^ in f.
|
||||
apply f.
|
||||
+ intros X1 X2 IH1 IH2 [H1 H2] a.
|
||||
simple refine (Trunc_ind _ _) ; cbn.
|
||||
intros [C1 | C2].
|
||||
++ apply (IH1 H1 a C1).
|
||||
++ apply (IH2 H2 a C2).
|
||||
Defined.
|
||||
|
||||
(** ** Extensionality proof *)
|
||||
|
||||
Lemma eq_subset' (X Y : FSet A) : X = Y <~> (U Y X = X) * (U X Y = Y).
|
||||
Proof.
|
||||
unshelve eapply BuildEquiv.
|
||||
{ intro H'. rewrite H'. split; apply union_idem. }
|
||||
unshelve esplit.
|
||||
{ intros [H1 H2]. etransitivity. apply H1^.
|
||||
rewrite comm. apply H2. }
|
||||
intro; apply path_prod; apply set_path2.
|
||||
all: intro; apply set_path2.
|
||||
Defined.
|
||||
|
||||
Lemma eq_subset (X Y : FSet A) :
|
||||
X = Y <~> (subset Y X * subset X Y).
|
||||
Proof.
|
||||
transitivity ((U Y X = X) * (U X Y = Y)).
|
||||
apply eq_subset'.
|
||||
symmetry.
|
||||
eapply equiv_functor_prod'; apply subset_union_equiv.
|
||||
Defined.
|
||||
|
||||
Theorem fset_ext (X Y : FSet A) :
|
||||
X = Y <~> (forall (a : A), isIn a X = isIn a Y).
|
||||
Proof.
|
||||
refine (@equiv_compose' _ _ _ _ _) ; [ | apply eq_subset ].
|
||||
refine (@equiv_compose' _ ((forall a, isIn a Y -> isIn a X)
|
||||
*(forall a, isIn a X -> isIn a Y)) _ _ _).
|
||||
- apply equiv_iff_hprop_uncurried.
|
||||
split.
|
||||
* intros [H1 H2 a].
|
||||
specialize (H1 a) ; specialize (H2 a).
|
||||
apply path_iff_hprop.
|
||||
apply H2.
|
||||
apply H1.
|
||||
* intros H1.
|
||||
split ; intro a ; intro H2.
|
||||
+ rewrite (H1 a).
|
||||
apply H2.
|
||||
+ rewrite <- (H1 a).
|
||||
apply H2.
|
||||
- eapply equiv_functor_prod' ;
|
||||
apply equiv_iff_hprop_uncurried ;
|
||||
split ; apply subset_isIn.
|
||||
Defined.
|
||||
|
||||
End ext.
|
||||
70
FiniteSets/fsets/monad.v
Normal file
70
FiniteSets/fsets/monad.v
Normal file
@@ -0,0 +1,70 @@
|
||||
(* [FSet] is a (strong and stable) finite powerset monad *)
|
||||
Require Import HoTT HitTactics.
|
||||
Require Export representations.definition fsets.properties.
|
||||
|
||||
Definition ffmap {A B : Type} : (A -> B) -> FSet A -> FSet B.
|
||||
Proof.
|
||||
intro f.
|
||||
hrecursion.
|
||||
- exact ∅.
|
||||
- intro a. exact {| f a |}.
|
||||
- exact U.
|
||||
- apply assoc.
|
||||
- apply comm.
|
||||
- apply nl.
|
||||
- apply nr.
|
||||
- simpl. intro x. apply idem.
|
||||
Defined.
|
||||
|
||||
Lemma ffmap_1 `{Funext} {A : Type} : @ffmap A A idmap = idmap.
|
||||
Proof.
|
||||
apply path_forall.
|
||||
intro x. hinduction x; try (intros; f_ap);
|
||||
try (intros; apply set_path2).
|
||||
Defined.
|
||||
|
||||
Global Instance fset_functorish `{Funext}: Functorish FSet
|
||||
:= { fmap := @ffmap; fmap_idmap := @ffmap_1 _ }.
|
||||
|
||||
Lemma ffmap_compose {A B C : Type} `{Funext} (f : A -> B) (g : B -> C) :
|
||||
fmap FSet (g o f) = fmap _ g o fmap _ f.
|
||||
Proof.
|
||||
apply path_forall. intro x.
|
||||
hrecursion x; try (intros; f_ap);
|
||||
try (intros; apply set_path2).
|
||||
Defined.
|
||||
|
||||
Definition join {A : Type} : FSet (FSet A) -> FSet A.
|
||||
Proof.
|
||||
hrecursion.
|
||||
- exact ∅.
|
||||
- exact idmap.
|
||||
- exact U.
|
||||
- apply assoc.
|
||||
- apply comm.
|
||||
- apply nl.
|
||||
- apply nr.
|
||||
- simpl. apply union_idem.
|
||||
Defined.
|
||||
|
||||
Lemma join_assoc {A : Type} (X : FSet (FSet (FSet A))) :
|
||||
join (ffmap join X) = join (join X).
|
||||
Proof.
|
||||
hrecursion X; try (intros; f_ap);
|
||||
try (intros; apply set_path2).
|
||||
Defined.
|
||||
|
||||
Lemma join_return_1 {A : Type} (X : FSet A) :
|
||||
join ({| X |}) = X.
|
||||
Proof. reflexivity. Defined.
|
||||
|
||||
Lemma join_return_fmap {A : Type} (X : FSet A) :
|
||||
join ({| X |}) = join (ffmap (fun x => {|x|}) X).
|
||||
Proof.
|
||||
hrecursion X; try (intros; f_ap);
|
||||
try (intros; apply set_path2).
|
||||
Defined.
|
||||
|
||||
Lemma join_fmap_return_1 {A : Type} (X : FSet A) :
|
||||
join (ffmap (fun x => {|x|}) X) = X.
|
||||
Proof. refine ((join_return_fmap _)^ @ join_return_1 _). Defined.
|
||||
88
FiniteSets/fsets/operations.v
Normal file
88
FiniteSets/fsets/operations.v
Normal file
@@ -0,0 +1,88 @@
|
||||
(* Operations on the [FSet A] for an arbitrary [A] *)
|
||||
Require Import HoTT HitTactics.
|
||||
Require Import representations.definition disjunction lattice.
|
||||
|
||||
Section operations.
|
||||
Context {A : Type}.
|
||||
Context `{Univalence}.
|
||||
|
||||
Definition isIn : A -> FSet A -> hProp.
|
||||
Proof.
|
||||
intros a.
|
||||
hrecursion.
|
||||
- exists Empty.
|
||||
exact _.
|
||||
- intro a'.
|
||||
exists (Trunc (-1) (a = a')).
|
||||
exact _.
|
||||
- apply lor.
|
||||
- intros ; apply lor_assoc. exact _.
|
||||
- intros ; apply lor_comm. exact _.
|
||||
- intros ; apply lor_nl. exact _.
|
||||
- intros ; apply lor_nr. exact _.
|
||||
- intros ; apply lor_idem. exact _.
|
||||
Defined.
|
||||
|
||||
Definition subset : FSet A -> FSet A -> hProp.
|
||||
Proof.
|
||||
intros X Y.
|
||||
hrecursion X.
|
||||
- exists Unit.
|
||||
exact _.
|
||||
- intros a.
|
||||
apply (isIn 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.
|
||||
|
||||
Definition comprehension :
|
||||
(A -> Bool) -> FSet A -> FSet A.
|
||||
Proof.
|
||||
intros P.
|
||||
hrecursion.
|
||||
- apply E.
|
||||
- intro a.
|
||||
refine (if (P a) then L a else E).
|
||||
- apply U.
|
||||
- apply assoc.
|
||||
- apply comm.
|
||||
- apply nl.
|
||||
- apply nr.
|
||||
- intros; simpl.
|
||||
destruct (P x).
|
||||
+ apply idem.
|
||||
+ apply nl.
|
||||
Defined.
|
||||
|
||||
Definition isEmpty :
|
||||
FSet A -> Bool.
|
||||
Proof.
|
||||
hrecursion.
|
||||
- apply true.
|
||||
- apply (fun _ => false).
|
||||
- apply andb.
|
||||
- intros. symmetry. eauto with bool_lattice_hints.
|
||||
- eauto with bool_lattice_hints.
|
||||
- eauto with bool_lattice_hints.
|
||||
- eauto with bool_lattice_hints.
|
||||
- eauto with bool_lattice_hints.
|
||||
Defined.
|
||||
|
||||
End operations.
|
||||
|
||||
Infix "∈" := isIn (at level 9, right associativity).
|
||||
Infix "⊆" := subset (at level 10, right associativity).
|
||||
41
FiniteSets/fsets/operations_decidable.v
Normal file
41
FiniteSets/fsets/operations_decidable.v
Normal 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.
|
||||
156
FiniteSets/fsets/properties.v
Normal file
156
FiniteSets/fsets/properties.v
Normal file
@@ -0,0 +1,156 @@
|
||||
Require Import HoTT HitTactics.
|
||||
Require Export representations.definition disjunction fsets.operations.
|
||||
|
||||
(* Lemmas relating operations to the membership predicate *)
|
||||
Section operations_isIn.
|
||||
|
||||
Context {A : Type}.
|
||||
Context `{Univalence}.
|
||||
|
||||
|
||||
|
||||
Lemma union_idem : forall x: FSet A, U x x = x.
|
||||
Proof.
|
||||
hinduction;
|
||||
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).
|
||||
f_ap.
|
||||
Defined.
|
||||
|
||||
(** ** Properties about subset relation. *)
|
||||
Lemma subset_union (X Y : FSet A) :
|
||||
subset X Y -> U X Y = Y.
|
||||
Proof.
|
||||
hinduction X; try (intros; apply path_forall; intro; apply set_path2).
|
||||
- intros. apply nl.
|
||||
- intros a. hinduction Y;
|
||||
try (intros; apply path_forall; intro; apply set_path2).
|
||||
+ intro.
|
||||
contradiction.
|
||||
+ intro a0.
|
||||
simple refine (Trunc_ind _ _).
|
||||
intro p ; cbn.
|
||||
rewrite p; apply idem.
|
||||
+ intros X1 X2 IH1 IH2.
|
||||
simple refine (Trunc_ind _ _).
|
||||
intros [e1 | e2].
|
||||
++ rewrite assoc.
|
||||
rewrite (IH1 e1).
|
||||
reflexivity.
|
||||
++ rewrite comm.
|
||||
rewrite <- assoc.
|
||||
rewrite (comm X2).
|
||||
rewrite (IH2 e2).
|
||||
reflexivity.
|
||||
- intros X1 X2 IH1 IH2 [G1 G2].
|
||||
rewrite <- assoc.
|
||||
rewrite (IH2 G2).
|
||||
apply (IH1 G1).
|
||||
Defined.
|
||||
|
||||
Lemma subset_union_l (X : FSet A) :
|
||||
forall Y, subset X (U X Y).
|
||||
Proof.
|
||||
hinduction X ;
|
||||
try (repeat (intro; intros; apply path_forall); intro; apply equiv_hprop_allpath ; apply _).
|
||||
- apply (fun _ => tt).
|
||||
- intros a Y.
|
||||
apply tr ; left ; apply tr ; reflexivity.
|
||||
- intros X1 X2 HX1 HX2 Y.
|
||||
split.
|
||||
* rewrite <- assoc. apply HX1.
|
||||
* rewrite (comm X1 X2). rewrite <- assoc. apply HX2.
|
||||
Defined.
|
||||
|
||||
(* Union and membership *)
|
||||
Lemma union_isIn (X Y : FSet A) (a : A) :
|
||||
isIn a (U X Y) = isIn a X ∨ isIn a Y.
|
||||
Proof.
|
||||
reflexivity.
|
||||
Defined.
|
||||
|
||||
Lemma comprehension_or : forall ϕ ψ (x: FSet A),
|
||||
comprehension (fun a => orb (ϕ a) (ψ a)) x = U (comprehension ϕ x)
|
||||
(comprehension ψ x).
|
||||
Proof.
|
||||
intros ϕ ψ.
|
||||
hinduction; try (intros; apply set_path2).
|
||||
- cbn. apply (union_idem _)^.
|
||||
- cbn. intros.
|
||||
destruct (ϕ a) ; destruct (ψ a) ; symmetry.
|
||||
* apply union_idem.
|
||||
* apply nr.
|
||||
* apply nl.
|
||||
* apply union_idem.
|
||||
- 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.
|
||||
|
||||
End operations_isIn.
|
||||
|
||||
(* Other properties *)
|
||||
Section properties.
|
||||
|
||||
Context {A : Type}.
|
||||
Context `{Univalence}.
|
||||
|
||||
(** isIn properties *)
|
||||
Lemma singleton_isIn (a b: A) : isIn a (L b) -> Trunc (-1) (a = b).
|
||||
Proof.
|
||||
apply idmap.
|
||||
Defined.
|
||||
|
||||
Lemma empty_isIn (a: A) : isIn a E -> Empty.
|
||||
Proof.
|
||||
apply idmap.
|
||||
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.
|
||||
|
||||
Lemma comprehension_subset : forall ϕ (X : FSet A),
|
||||
U (comprehension ϕ X) X = X.
|
||||
Proof.
|
||||
intros ϕ.
|
||||
hrecursion; try (intros ; apply set_path2) ; cbn.
|
||||
- apply union_idem.
|
||||
- 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).
|
||||
rewrite assoc.
|
||||
rewrite P.
|
||||
rewrite <- assoc.
|
||||
rewrite Q.
|
||||
reflexivity.
|
||||
Defined.
|
||||
|
||||
End properties.
|
||||
277
FiniteSets/fsets/properties_decidable.v
Normal file
277
FiniteSets/fsets/properties_decidable.v
Normal file
@@ -0,0 +1,277 @@
|
||||
(* Properties of [FSet A] where [A] has decidable equality *)
|
||||
Require Import HoTT HitTactics.
|
||||
From fsets Require Export properties extensionality operations_decidable.
|
||||
Require Export lattice.
|
||||
|
||||
(* Lemmas relating operations to the membership predicate *)
|
||||
Section operations_isIn.
|
||||
|
||||
Context {A : Type} `{DecidablePaths A} `{Univalence}.
|
||||
|
||||
Lemma ext : forall (S T : FSet A), (forall a, isIn_b a S = isIn_b a T) -> S = T.
|
||||
Proof.
|
||||
intros X Y H2.
|
||||
apply fset_ext.
|
||||
intro a.
|
||||
specialize (H2 a).
|
||||
unfold isIn_b, dec in H2.
|
||||
destruct (isIn_decidable a X) ; destruct (isIn_decidable a Y).
|
||||
- apply path_iff_hprop ; intro ; assumption.
|
||||
- contradiction (true_ne_false).
|
||||
- contradiction (true_ne_false) ; apply H2^.
|
||||
- apply path_iff_hprop ; intro ; contradiction.
|
||||
Defined.
|
||||
|
||||
Lemma empty_isIn (a : A) :
|
||||
isIn_b a ∅ = false.
|
||||
Proof.
|
||||
reflexivity.
|
||||
Defined.
|
||||
|
||||
Lemma L_isIn (a b : A) :
|
||||
isIn a {|b|} -> a = b.
|
||||
Proof.
|
||||
intros. strip_truncations. assumption.
|
||||
Defined.
|
||||
|
||||
Lemma L_isIn_b_true (a b : A) (p : a = b) :
|
||||
isIn_b a {|b|} = true.
|
||||
Proof.
|
||||
unfold isIn_b, dec.
|
||||
destruct (isIn_decidable a {|b|}) as [n | n] .
|
||||
- reflexivity.
|
||||
- simpl in n.
|
||||
contradiction (n (tr p)).
|
||||
Defined.
|
||||
|
||||
Lemma L_isIn_b_aa (a : A) :
|
||||
isIn_b a {|a|} = true.
|
||||
Proof.
|
||||
apply L_isIn_b_true ; reflexivity.
|
||||
Defined.
|
||||
|
||||
Lemma L_isIn_b_false (a b : A) (p : a <> b) :
|
||||
isIn_b a {|b|} = false.
|
||||
Proof.
|
||||
unfold isIn_b, dec.
|
||||
destruct (isIn_decidable a {|b|}).
|
||||
- simpl in t.
|
||||
strip_truncations.
|
||||
contradiction.
|
||||
- reflexivity.
|
||||
Defined.
|
||||
|
||||
(* Union and membership *)
|
||||
Lemma union_isIn (X Y : FSet A) (a : A) :
|
||||
isIn_b a (U X Y) = orb (isIn_b a X) (isIn_b a Y).
|
||||
Proof.
|
||||
unfold isIn_b ; unfold dec.
|
||||
simpl.
|
||||
destruct (isIn_decidable a X) ; destruct (isIn_decidable a Y) ; reflexivity.
|
||||
Defined.
|
||||
|
||||
Lemma intersection_isIn (X Y: FSet A) (a : A) :
|
||||
isIn_b a (intersection X Y) = andb (isIn_b a X) (isIn_b a Y).
|
||||
Proof.
|
||||
hinduction X; try (intros ; apply set_path2).
|
||||
- reflexivity.
|
||||
- intro b.
|
||||
destruct (dec (a = b)).
|
||||
* rewrite p.
|
||||
destruct (isIn_b b Y) ; symmetry ; eauto with bool_lattice_hints.
|
||||
* destruct (isIn_b b Y) ; destruct (isIn_b a Y) ; symmetry ; eauto with bool_lattice_hints.
|
||||
+ rewrite and_false.
|
||||
symmetry.
|
||||
apply (L_isIn_b_false a b n).
|
||||
+ rewrite and_true.
|
||||
apply (L_isIn_b_false a b n).
|
||||
- intros X1 X2 P Q.
|
||||
rewrite union_isIn ; rewrite union_isIn.
|
||||
rewrite P.
|
||||
rewrite Q.
|
||||
unfold isIn_b, dec.
|
||||
destruct (isIn_decidable a X1)
|
||||
; destruct (isIn_decidable a X2)
|
||||
; destruct (isIn_decidable a Y)
|
||||
; reflexivity.
|
||||
Defined.
|
||||
|
||||
Lemma comprehension_isIn (Y : FSet A) (ϕ : A -> Bool) (a : A) :
|
||||
isIn_b a (comprehension ϕ Y) = andb (isIn_b a Y) (ϕ a).
|
||||
Proof.
|
||||
hinduction Y ; try (intros; apply set_path2).
|
||||
- apply empty_isIn.
|
||||
- intro b.
|
||||
destruct (isIn_decidable a {|b|}).
|
||||
* simpl in t.
|
||||
strip_truncations.
|
||||
rewrite t.
|
||||
destruct (ϕ b).
|
||||
** rewrite (L_isIn_b_true _ _ idpath).
|
||||
eauto with bool_lattice_hints.
|
||||
** rewrite empty_isIn ; rewrite (L_isIn_b_true _ _ idpath).
|
||||
eauto with bool_lattice_hints.
|
||||
* destruct (ϕ b).
|
||||
** rewrite L_isIn_b_false.
|
||||
*** eauto with bool_lattice_hints.
|
||||
*** intro.
|
||||
apply (n (tr X)).
|
||||
** rewrite empty_isIn.
|
||||
rewrite L_isIn_b_false.
|
||||
*** eauto with bool_lattice_hints.
|
||||
*** intro.
|
||||
apply (n (tr X)).
|
||||
- intros.
|
||||
Opaque isIn_b.
|
||||
rewrite ?union_isIn.
|
||||
rewrite X.
|
||||
rewrite X0.
|
||||
assert (forall b1 b2 b3,
|
||||
(b1 && b2 || b3 && b2)%Bool = ((b1 || b3) && b2)%Bool).
|
||||
{ intros ; destruct b1, b2, b3 ; reflexivity. }
|
||||
apply X1.
|
||||
Defined.
|
||||
End operations_isIn.
|
||||
|
||||
Global Opaque isIn_b.
|
||||
(* Some suporting tactics *)
|
||||
Ltac simplify_isIn :=
|
||||
repeat (rewrite union_isIn
|
||||
|| rewrite L_isIn_b_aa
|
||||
|| rewrite intersection_isIn
|
||||
|| rewrite comprehension_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 `{Univalence}.
|
||||
|
||||
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 intersection (@U A) (@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.
|
||||
|
||||
(* Comprehension properties *)
|
||||
Section comprehension_properties.
|
||||
|
||||
Opaque isIn_b.
|
||||
|
||||
Context {A : Type}.
|
||||
Context {A_deceq : DecidablePaths A}.
|
||||
Context `{Univalence}.
|
||||
|
||||
Lemma comprehension_or : forall ϕ ψ (x: FSet A),
|
||||
comprehension (fun a => orb (ϕ a) (ψ a)) x
|
||||
= U (comprehension ϕ x) (comprehension ψ x).
|
||||
Proof.
|
||||
toBool.
|
||||
Defined.
|
||||
|
||||
(** comprehension properties *)
|
||||
Lemma comprehension_false Y : comprehension (fun (_ : A) => false) Y = E.
|
||||
Proof.
|
||||
toBool.
|
||||
Defined.
|
||||
|
||||
Lemma comprehension_all : forall (X : FSet A),
|
||||
comprehension (fun a => isIn_b a X) X = X.
|
||||
Proof.
|
||||
toBool.
|
||||
Defined.
|
||||
|
||||
Lemma comprehension_subset : forall ϕ (X : FSet A),
|
||||
U (comprehension ϕ X) X = X.
|
||||
Proof.
|
||||
toBool.
|
||||
Defined.
|
||||
|
||||
End comprehension_properties.
|
||||
|
||||
(* With extensionality we can prove decidable equality *)
|
||||
Section dec_eq.
|
||||
Context (A : Type) `{DecidablePaths A} `{Univalence}.
|
||||
|
||||
Instance decidable_prod {X Y : Type} `{Decidable X} `{Decidable Y} :
|
||||
Decidable (X * Y).
|
||||
Proof.
|
||||
unfold Decidable in *.
|
||||
destruct H1 as [x | nx] ; destruct H2 as [y | ny].
|
||||
- left ; split ; assumption.
|
||||
- right. intros [p1 p2]. contradiction.
|
||||
- right. intros [p1 p2]. contradiction.
|
||||
- right. intros [p1 p2]. contradiction.
|
||||
Defined.
|
||||
|
||||
Instance fsets_dec_eq : DecidablePaths (FSet A).
|
||||
Proof.
|
||||
intros X Y.
|
||||
apply (decidable_equiv' ((subset Y X) * (subset X Y)) (eq_subset X Y)^-1).
|
||||
apply _.
|
||||
Defined.
|
||||
|
||||
End dec_eq.
|
||||
Reference in New Issue
Block a user