mirror of
https://github.com/nmvdw/HITs-Examples
synced 2025-12-16 23:23:50 +01:00
Improved notatio
This commit is contained in:
@@ -9,17 +9,23 @@ Section characterize_isIn.
|
||||
Context `{Univalence}.
|
||||
|
||||
(** isIn properties *)
|
||||
Definition empty_isIn (a: A) : a ∈ E -> Empty := idmap.
|
||||
Definition empty_isIn (a: A) : a ∈ ∅ -> Empty := idmap.
|
||||
|
||||
Definition singleton_isIn (a b: A) : a ∈ {|b|} -> Trunc (-1) (a = b) := idmap.
|
||||
|
||||
Definition union_isIn (X Y : FSet A) (a : A)
|
||||
: a ∈ X ∪ Y = a ∈ X ∨ a ∈ Y := idpath.
|
||||
|
||||
Lemma comprehension_isIn (ϕ : A -> Bool) (a : A) : forall X : FSet A,
|
||||
a ∈ (comprehension ϕ X) = if ϕ a then a ∈ X else False_hp.
|
||||
Lemma comprehension_union (ϕ : A -> Bool) : forall X Y : FSet A,
|
||||
{|X ∪ Y & ϕ|} = {|X & ϕ|} ∪ {|Y & ϕ|}.
|
||||
Proof.
|
||||
hinduction ; try (intros ; apply set_path2) ; cbn.
|
||||
reflexivity.
|
||||
Defined.
|
||||
|
||||
Lemma comprehension_isIn (ϕ : A -> Bool) (a : A) : forall X : FSet A,
|
||||
a ∈ {|X & ϕ|} = if ϕ a then a ∈ X else False_hp.
|
||||
Proof.
|
||||
hinduction ; try (intros ; apply set_path2).
|
||||
- destruct (ϕ a) ; reflexivity.
|
||||
- intros b.
|
||||
assert (forall c d, ϕ a = c -> ϕ b = d ->
|
||||
@@ -36,6 +42,8 @@ Section characterize_isIn.
|
||||
}
|
||||
apply (X (ϕ a) (ϕ b) idpath idpath).
|
||||
- intros X Y H1 H2.
|
||||
rewrite comprehension_union.
|
||||
rewrite union_isIn.
|
||||
rewrite H1, H2.
|
||||
destruct (ϕ a).
|
||||
* reflexivity.
|
||||
@@ -44,11 +52,14 @@ Section characterize_isIn.
|
||||
destruct Z ; assumption.
|
||||
** intros ; apply tr ; right ; assumption.
|
||||
Defined.
|
||||
End characterize_isIn.
|
||||
|
||||
Context {B : Type}.
|
||||
|
||||
Section product_isIn.
|
||||
Context {A B : Type}.
|
||||
Context `{Univalence}.
|
||||
|
||||
Lemma isIn_singleproduct (a : A) (b : B) (c : A) : forall (Y : FSet B),
|
||||
isIn (a, b) (single_product c Y) = land (BuildhProp (Trunc (-1) (a = c))) (isIn b Y).
|
||||
(a, b) ∈ (single_product c Y) = land (BuildhProp (Trunc (-1) (a = c))) (b ∈ Y).
|
||||
Proof.
|
||||
hinduction ; try (intros ; apply path_ishprop).
|
||||
- apply path_hprop ; symmetry ; apply prod_empty_r.
|
||||
@@ -62,16 +73,16 @@ Section characterize_isIn.
|
||||
rewrite Z1, Z2.
|
||||
apply (tr idpath).
|
||||
- intros X1 X2 HX1 HX2.
|
||||
apply path_iff_hprop.
|
||||
* intros X.
|
||||
strip_truncations.
|
||||
destruct X as [H1 | H1] ; rewrite ?HX1, ?HX2 in H1 ; destruct H1 as [H1 H2].
|
||||
** split.
|
||||
*** apply H1.
|
||||
*** apply (tr(inl H2)).
|
||||
** split.
|
||||
*** apply H1.
|
||||
*** apply (tr(inr H2)).
|
||||
apply path_iff_hprop ; rewrite ?union_isIn.
|
||||
* intros X.
|
||||
strip_truncations.
|
||||
destruct X as [H1 | H1] ; rewrite ?HX1, ?HX2 in H1 ; destruct H1 as [H1 H2].
|
||||
** split.
|
||||
*** apply H1.
|
||||
*** apply (tr(inl H2)).
|
||||
** split.
|
||||
*** apply H1.
|
||||
*** apply (tr(inr H2)).
|
||||
* intros [H1 H2].
|
||||
strip_truncations.
|
||||
apply tr.
|
||||
@@ -84,15 +95,16 @@ Section characterize_isIn.
|
||||
Defined.
|
||||
|
||||
Definition isIn_product (a : A) (b : B) (X : FSet A) (Y : FSet B) :
|
||||
isIn (a,b) (product X Y) = land (isIn a X) (isIn b Y).
|
||||
(a,b) ∈ (product X Y) = land (a ∈ X) (b ∈ Y).
|
||||
Proof.
|
||||
hinduction X ; try (intros ; apply path_ishprop).
|
||||
- apply path_hprop ; symmetry ; apply prod_empty_l.
|
||||
- intros.
|
||||
apply isIn_singleproduct.
|
||||
- intros X1 X2 HX1 HX2.
|
||||
rewrite (union_isIn (product X1 Y)).
|
||||
rewrite HX1, HX2.
|
||||
apply path_iff_hprop.
|
||||
apply path_iff_hprop ; rewrite ?union_isIn.
|
||||
* intros X.
|
||||
strip_truncations.
|
||||
destruct X as [[H3 H4] | [H3 H4]] ; split ; try (apply H4).
|
||||
@@ -104,7 +116,7 @@ Section characterize_isIn.
|
||||
** left ; split ; assumption.
|
||||
** right ; split ; assumption.
|
||||
Defined.
|
||||
End characterize_isIn.
|
||||
End product_isIn.
|
||||
|
||||
Ltac simplify_isIn :=
|
||||
repeat (rewrite union_isIn
|
||||
@@ -120,8 +132,17 @@ Section properties.
|
||||
Context {A : Type}.
|
||||
Context `{Univalence}.
|
||||
|
||||
Instance: bottom(FSet A) := ∅.
|
||||
Instance: maximum(FSet A) := U.
|
||||
Instance: bottom(FSet A).
|
||||
Proof.
|
||||
unfold bottom.
|
||||
apply ∅.
|
||||
Defined.
|
||||
|
||||
Instance: maximum(FSet A).
|
||||
Proof.
|
||||
intros x y.
|
||||
apply (x ∪ y).
|
||||
Defined.
|
||||
|
||||
Global Instance joinsemilattice_fset : JoinSemiLattice (FSet A).
|
||||
Proof.
|
||||
@@ -129,26 +150,32 @@ Section properties.
|
||||
Defined.
|
||||
|
||||
(** comprehension properties *)
|
||||
Lemma comprehension_false Y : comprehension (fun (_ : A) => false) Y = ∅.
|
||||
Lemma comprehension_false X : {|X & (fun (_ : A) => false)|} = ∅.
|
||||
Proof.
|
||||
toHProp.
|
||||
Defined.
|
||||
|
||||
Lemma comprehension_subset : forall ϕ (X : FSet A),
|
||||
(comprehension ϕ X) ∪ X = X.
|
||||
{|X & ϕ|} ∪ X = X.
|
||||
Proof.
|
||||
toHProp.
|
||||
destruct (ϕ a) ; eauto with lattice_hints typeclass_instances.
|
||||
Defined.
|
||||
|
||||
Lemma comprehension_or : forall ϕ ψ (x: FSet A),
|
||||
comprehension (fun a => orb (ϕ a) (ψ a)) x
|
||||
= (comprehension ϕ x) ∪ (comprehension ψ x).
|
||||
Lemma comprehension_or : forall ϕ ψ (X : FSet A),
|
||||
{|X & (fun a => orb (ϕ a) (ψ a))|}
|
||||
= {|X & ϕ|} ∪ {|X & ψ|}.
|
||||
Proof.
|
||||
toHProp.
|
||||
symmetry ; destruct (ϕ a) ; destruct (ψ a)
|
||||
; eauto with lattice_hints typeclass_instances.
|
||||
Defined.
|
||||
|
||||
Lemma comprehension_all : forall (X : FSet A),
|
||||
comprehension (fun a => true) X = X.
|
||||
Proof.
|
||||
toHProp.
|
||||
Defined.
|
||||
|
||||
Lemma merely_choice : forall X : FSet A, hor (X = ∅) (hexists (fun a => a ∈ X)).
|
||||
Proof.
|
||||
@@ -161,7 +188,7 @@ Section properties.
|
||||
destruct TX as [XE | HX] ; destruct TY as [YE | HY] ; try(strip_truncations ; apply tr).
|
||||
* refine (tr (inl _)).
|
||||
rewrite XE, YE.
|
||||
apply (union_idem E).
|
||||
apply (union_idem ∅).
|
||||
* destruct HY as [a Ya].
|
||||
refine (inr (tr _)).
|
||||
exists a.
|
||||
|
||||
Reference in New Issue
Block a user