mirror of https://github.com/nmvdw/HITs-Examples
Some cleaning in notation
This commit is contained in:
parent
1bab2206a3
commit
1e373364b2
|
@ -8,7 +8,7 @@ Section ext.
|
||||||
Context `{Univalence}.
|
Context `{Univalence}.
|
||||||
|
|
||||||
Lemma subset_union_equiv
|
Lemma subset_union_equiv
|
||||||
: forall X Y : FSet A, subset X Y <~> U X Y = Y.
|
: forall X Y : FSet A, X ⊆ Y <~> X ∪ Y = Y.
|
||||||
Proof.
|
Proof.
|
||||||
intros X Y.
|
intros X Y.
|
||||||
eapply equiv_iff_hprop_uncurried.
|
eapply equiv_iff_hprop_uncurried.
|
||||||
|
@ -20,8 +20,8 @@ Section ext.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Lemma subset_isIn (X Y : FSet A) :
|
Lemma subset_isIn (X Y : FSet A) :
|
||||||
(forall (a : A), isIn a X -> isIn a Y)
|
(forall (a : A), a ∈ X -> a ∈ Y)
|
||||||
<~> (subset X Y).
|
<~> (X ⊆ Y).
|
||||||
Proof.
|
Proof.
|
||||||
eapply equiv_iff_hprop_uncurried.
|
eapply equiv_iff_hprop_uncurried.
|
||||||
split.
|
split.
|
||||||
|
@ -32,19 +32,17 @@ Section ext.
|
||||||
apply f.
|
apply f.
|
||||||
apply tr ; reflexivity.
|
apply tr ; reflexivity.
|
||||||
+ intros X1 X2 H1 H2 f.
|
+ intros X1 X2 H1 H2 f.
|
||||||
enough (subset X1 Y).
|
enough (X1 ⊆ Y).
|
||||||
enough (subset X2 Y).
|
enough (X2 ⊆ Y).
|
||||||
{ split. apply X. apply X0. }
|
{ split. apply X. apply X0. }
|
||||||
* apply H2.
|
* apply H2.
|
||||||
intros a Ha.
|
intros a Ha.
|
||||||
apply f.
|
refine (f _ (tr _)).
|
||||||
apply tr.
|
|
||||||
right.
|
right.
|
||||||
apply Ha.
|
apply Ha.
|
||||||
* apply H1.
|
* apply H1.
|
||||||
intros a Ha.
|
intros a Ha.
|
||||||
apply f.
|
refine (f _ (tr _)).
|
||||||
apply tr.
|
|
||||||
left.
|
left.
|
||||||
apply Ha.
|
apply Ha.
|
||||||
- hinduction X ;
|
- hinduction X ;
|
||||||
|
@ -64,7 +62,7 @@ Section ext.
|
||||||
|
|
||||||
(** ** Extensionality proof *)
|
(** ** Extensionality proof *)
|
||||||
|
|
||||||
Lemma eq_subset' (X Y : FSet A) : X = Y <~> (U Y X = X) * (U X Y = Y).
|
Lemma eq_subset' (X Y : FSet A) : X = Y <~> (Y ∪ X = X) * (X ∪ Y = Y).
|
||||||
Proof.
|
Proof.
|
||||||
unshelve eapply BuildEquiv.
|
unshelve eapply BuildEquiv.
|
||||||
{ intro H'. rewrite H'. split; apply union_idem. }
|
{ intro H'. rewrite H'. split; apply union_idem. }
|
||||||
|
@ -76,20 +74,20 @@ Section ext.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Lemma eq_subset (X Y : FSet A) :
|
Lemma eq_subset (X Y : FSet A) :
|
||||||
X = Y <~> (subset Y X * subset X Y).
|
X = Y <~> (Y ⊆ X * X ⊆ Y).
|
||||||
Proof.
|
Proof.
|
||||||
transitivity ((U Y X = X) * (U X Y = Y)).
|
transitivity ((Y ∪ X = X) * (X ∪ Y = Y)).
|
||||||
apply eq_subset'.
|
apply eq_subset'.
|
||||||
symmetry.
|
symmetry.
|
||||||
eapply equiv_functor_prod'; apply subset_union_equiv.
|
eapply equiv_functor_prod'; apply subset_union_equiv.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Theorem fset_ext (X Y : FSet A) :
|
Theorem fset_ext (X Y : FSet A) :
|
||||||
X = Y <~> (forall (a : A), isIn a X = isIn a Y).
|
X = Y <~> (forall (a : A), a ∈ X = a ∈ Y).
|
||||||
Proof.
|
Proof.
|
||||||
refine (@equiv_compose' _ _ _ _ _) ; [ | apply eq_subset ].
|
refine (@equiv_compose' _ _ _ _ _) ; [ | apply eq_subset ].
|
||||||
refine (@equiv_compose' _ ((forall a, isIn a Y -> isIn a X)
|
refine (@equiv_compose' _ ((forall a, a ∈ Y -> a ∈ X)
|
||||||
*(forall a, isIn a X -> isIn a Y)) _ _ _).
|
*(forall a, a ∈ X -> a ∈ Y)) _ _ _).
|
||||||
- apply equiv_iff_hprop_uncurried.
|
- apply equiv_iff_hprop_uncurried.
|
||||||
split.
|
split.
|
||||||
* intros [H1 H2 a].
|
* intros [H1 H2 a].
|
||||||
|
|
|
@ -11,7 +11,7 @@ Section Length.
|
||||||
|
|
||||||
Opaque isIn_b.
|
Opaque isIn_b.
|
||||||
|
|
||||||
Definition length (x: FSetC A) : nat.
|
Definition length (x : FSetC A) : nat.
|
||||||
Proof.
|
Proof.
|
||||||
simple refine (FSetC_ind A _ _ _ _ _ _ x ); simpl.
|
simple refine (FSetC_ind A _ _ _ _ _ _ x ); simpl.
|
||||||
- exact 0.
|
- exact 0.
|
||||||
|
@ -31,7 +31,7 @@ Section Length.
|
||||||
|
|
||||||
Definition length_FSet (x: FSet A) := length (FSet_to_FSetC x).
|
Definition length_FSet (x: FSet A) := length (FSet_to_FSetC x).
|
||||||
|
|
||||||
Lemma length_singleton: forall (a: A), length_FSet (L a) = 1.
|
Lemma length_singleton: forall (a: A), length_FSet ({|a|}) = 1.
|
||||||
Proof.
|
Proof.
|
||||||
intro a.
|
intro a.
|
||||||
cbn. reflexivity.
|
cbn. reflexivity.
|
||||||
|
|
|
@ -3,84 +3,78 @@ Require Import HoTT HitTactics.
|
||||||
Require Import representations.definition disjunction lattice.
|
Require Import representations.definition disjunction lattice.
|
||||||
|
|
||||||
Section operations.
|
Section operations.
|
||||||
Context {A : Type}.
|
Context {A : Type}.
|
||||||
Context `{Univalence}.
|
Context `{Univalence}.
|
||||||
|
|
||||||
Definition isIn : A -> FSet A -> hProp.
|
Definition isIn : A -> FSet A -> hProp.
|
||||||
Proof.
|
Proof.
|
||||||
intros a.
|
intros a.
|
||||||
hrecursion.
|
hrecursion.
|
||||||
- exists Empty.
|
- exists Empty.
|
||||||
exact _.
|
exact _.
|
||||||
- intro a'.
|
- intro a'.
|
||||||
exists (Trunc (-1) (a = a')).
|
exists (Trunc (-1) (a = a')).
|
||||||
exact _.
|
exact _.
|
||||||
- apply lor.
|
- apply lor.
|
||||||
- intros ; symmetry ; apply lor_assoc.
|
- intros ; symmetry ; apply lor_assoc.
|
||||||
- apply lor_commutative.
|
- apply lor_commutative.
|
||||||
- apply lor_nl.
|
- apply lor_nl.
|
||||||
- apply lor_nr.
|
- apply lor_nr.
|
||||||
- intros ; apply lor_idem.
|
- intros ; apply lor_idem.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Definition subset : FSet A -> FSet A -> hProp.
|
Definition subset : FSet A -> FSet A -> hProp.
|
||||||
Proof.
|
Proof.
|
||||||
intros X Y.
|
intros X Y.
|
||||||
hrecursion X.
|
hrecursion X.
|
||||||
- exists Unit.
|
- exists Unit.
|
||||||
exact _.
|
exact _.
|
||||||
- intros a.
|
- intros a.
|
||||||
apply (isIn a Y).
|
apply (isIn a Y).
|
||||||
- intros X1 X2.
|
- intros X1 X2.
|
||||||
exists (prod X1 X2).
|
exists (prod X1 X2).
|
||||||
exact _.
|
exact _.
|
||||||
- intros.
|
- intros.
|
||||||
apply path_trunctype ; apply equiv_prod_assoc.
|
apply path_trunctype ; apply equiv_prod_assoc.
|
||||||
- intros.
|
- intros.
|
||||||
apply path_trunctype ; apply equiv_prod_symm.
|
apply path_trunctype ; apply equiv_prod_symm.
|
||||||
- intros.
|
- intros.
|
||||||
apply path_trunctype ; apply prod_unit_l.
|
apply path_trunctype ; apply prod_unit_l.
|
||||||
- intros.
|
- intros.
|
||||||
apply path_trunctype ; apply prod_unit_r.
|
apply path_trunctype ; apply prod_unit_r.
|
||||||
- intros a'.
|
- intros a'.
|
||||||
apply path_iff_hprop ; cbn.
|
apply path_iff_hprop ; cbn.
|
||||||
* intros [p1 p2]. apply p1.
|
* intros [p1 p2]. apply p1.
|
||||||
* intros p.
|
* intros p.
|
||||||
split ; apply p.
|
split ; apply p.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Definition comprehension :
|
Definition comprehension :
|
||||||
(A -> Bool) -> FSet A -> FSet A.
|
(A -> Bool) -> FSet A -> FSet A.
|
||||||
Proof.
|
Proof.
|
||||||
intros P.
|
intros P.
|
||||||
hrecursion.
|
hrecursion.
|
||||||
- apply E.
|
- apply ∅.
|
||||||
- intro a.
|
- intro a.
|
||||||
refine (if (P a) then L a else E).
|
refine (if (P a) then {|a|} else ∅).
|
||||||
- apply U.
|
- apply U.
|
||||||
- apply assoc.
|
- apply assoc.
|
||||||
- apply comm.
|
- apply comm.
|
||||||
- apply nl.
|
- apply nl.
|
||||||
- apply nr.
|
- apply nr.
|
||||||
- intros; simpl.
|
- intros; simpl.
|
||||||
destruct (P x).
|
destruct (P x).
|
||||||
+ apply idem.
|
+ apply idem.
|
||||||
+ apply nl.
|
+ apply nl.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Definition isEmpty :
|
Definition isEmpty :
|
||||||
FSet A -> Bool.
|
FSet A -> Bool.
|
||||||
Proof.
|
Proof.
|
||||||
hrecursion.
|
simple refine (FSet_rec _ _ _ true (fun _ => false) andb _ _ _ _ _)
|
||||||
- apply true.
|
; try eauto with bool_lattice_hints typeclass_instances.
|
||||||
- apply (fun _ => false).
|
intros ; symmetry ; eauto with lattice_hints typeclass_instances.
|
||||||
- apply andb.
|
Defined.
|
||||||
- intros. symmetry. eauto with lattice_hints typeclass_instances.
|
|
||||||
- eauto with bool_lattice_hints typeclass_instances.
|
|
||||||
- eauto with bool_lattice_hints typeclass_instances.
|
|
||||||
- eauto with bool_lattice_hints typeclass_instances.
|
|
||||||
- eauto with bool_lattice_hints typeclass_instances.
|
|
||||||
Defined.
|
|
||||||
|
|
||||||
End operations.
|
End operations.
|
||||||
|
|
||||||
|
|
|
@ -7,7 +7,7 @@ Section operations_isIn.
|
||||||
Context {A : Type}.
|
Context {A : Type}.
|
||||||
Context `{Univalence}.
|
Context `{Univalence}.
|
||||||
|
|
||||||
Lemma union_idem : forall x: FSet A, U x x = x.
|
Lemma union_idem : forall x: FSet A, x ∪ x = x.
|
||||||
Proof.
|
Proof.
|
||||||
hinduction ; try (intros ; apply set_path2).
|
hinduction ; try (intros ; apply set_path2).
|
||||||
- apply nl.
|
- apply nl.
|
||||||
|
@ -24,7 +24,7 @@ Section operations_isIn.
|
||||||
|
|
||||||
(** ** Properties about subset relation. *)
|
(** ** Properties about subset relation. *)
|
||||||
Lemma subset_union (X Y : FSet A) :
|
Lemma subset_union (X Y : FSet A) :
|
||||||
subset X Y -> U X Y = Y.
|
X ⊆ Y -> X ∪ Y = Y.
|
||||||
Proof.
|
Proof.
|
||||||
hinduction X ; try (intros; apply path_forall; intro; apply set_path2).
|
hinduction X ; try (intros; apply path_forall; intro; apply set_path2).
|
||||||
- intros. apply nl.
|
- intros. apply nl.
|
||||||
|
@ -54,7 +54,7 @@ Section operations_isIn.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Lemma subset_union_l (X : FSet A) :
|
Lemma subset_union_l (X : FSet A) :
|
||||||
forall Y, subset X (U X Y).
|
forall Y, X ⊆ X ∪ Y.
|
||||||
Proof.
|
Proof.
|
||||||
hinduction X ; try (repeat (intro; intros; apply path_forall);
|
hinduction X ; try (repeat (intro; intros; apply path_forall);
|
||||||
intro ; apply equiv_hprop_allpath ; apply _).
|
intro ; apply equiv_hprop_allpath ; apply _).
|
||||||
|
@ -69,8 +69,8 @@ Section operations_isIn.
|
||||||
|
|
||||||
(* simplify it using extensionality *)
|
(* simplify it using extensionality *)
|
||||||
Lemma comprehension_or : forall ϕ ψ (x: FSet A),
|
Lemma comprehension_or : forall ϕ ψ (x: FSet A),
|
||||||
comprehension (fun a => orb (ϕ a) (ψ a)) x = U (comprehension ϕ x)
|
comprehension (fun a => orb (ϕ a) (ψ a)) x
|
||||||
(comprehension ψ x).
|
= (comprehension ϕ x) ∪ (comprehension ψ x).
|
||||||
Proof.
|
Proof.
|
||||||
intros ϕ ψ.
|
intros ϕ ψ.
|
||||||
hinduction ; try (intros; apply set_path2).
|
hinduction ; try (intros; apply set_path2).
|
||||||
|
@ -101,15 +101,15 @@ Section properties.
|
||||||
Context `{Univalence}.
|
Context `{Univalence}.
|
||||||
|
|
||||||
(** isIn properties *)
|
(** isIn properties *)
|
||||||
Definition empty_isIn (a: A) : isIn a E -> Empty := idmap.
|
Definition empty_isIn (a: A) : a ∈ E -> Empty := idmap.
|
||||||
|
|
||||||
Definition singleton_isIn (a b: A) : isIn a (L b) -> Trunc (-1) (a = b) := idmap.
|
Definition singleton_isIn (a b: A) : a ∈ {|b|} -> Trunc (-1) (a = b) := idmap.
|
||||||
|
|
||||||
Definition union_isIn (X Y : FSet A) (a : A)
|
Definition union_isIn (X Y : FSet A) (a : A)
|
||||||
: isIn a (U X Y) = isIn a X ∨ isIn a Y := idpath.
|
: a ∈ X ∪ Y = a ∈ X ∨ a ∈ Y := idpath.
|
||||||
|
|
||||||
Lemma comprehension_isIn (ϕ : A -> Bool) (a : A) : forall X : FSet A,
|
Lemma comprehension_isIn (ϕ : A -> Bool) (a : A) : forall X : FSet A,
|
||||||
isIn a (comprehension ϕ X) = if ϕ a then isIn a X else False_hp.
|
a ∈ (comprehension ϕ X) = if ϕ a then a ∈ X else False_hp.
|
||||||
Proof.
|
Proof.
|
||||||
hinduction ; try (intros ; apply set_path2) ; cbn.
|
hinduction ; try (intros ; apply set_path2) ; cbn.
|
||||||
- destruct (ϕ a) ; reflexivity.
|
- destruct (ϕ a) ; reflexivity.
|
||||||
|
@ -139,7 +139,7 @@ Section properties.
|
||||||
|
|
||||||
(* The proof can be simplified using extensionality *)
|
(* The proof can be simplified using extensionality *)
|
||||||
(** comprehension properties *)
|
(** comprehension properties *)
|
||||||
Lemma comprehension_false Y : comprehension (fun (_ : A) => false) Y = E.
|
Lemma comprehension_false Y : comprehension (fun (_ : A) => false) Y = ∅.
|
||||||
Proof.
|
Proof.
|
||||||
hrecursion Y; try (intros; apply set_path2).
|
hrecursion Y; try (intros; apply set_path2).
|
||||||
- reflexivity.
|
- reflexivity.
|
||||||
|
@ -151,7 +151,7 @@ Section properties.
|
||||||
|
|
||||||
(* Can be simplified using extensionality *)
|
(* Can be simplified using extensionality *)
|
||||||
Lemma comprehension_subset : forall ϕ (X : FSet A),
|
Lemma comprehension_subset : forall ϕ (X : FSet A),
|
||||||
U (comprehension ϕ X) X = X.
|
(comprehension ϕ X) ∪ X = X.
|
||||||
Proof.
|
Proof.
|
||||||
intros ϕ.
|
intros ϕ.
|
||||||
hrecursion; try (intros ; apply set_path2) ; cbn.
|
hrecursion; try (intros ; apply set_path2) ; cbn.
|
||||||
|
@ -171,7 +171,7 @@ Section properties.
|
||||||
reflexivity.
|
reflexivity.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Lemma merely_choice : forall X : FSet A, hor (X = E) (hexists (fun a => isIn a X)).
|
Lemma merely_choice : forall X : FSet A, hor (X = ∅) (hexists (fun a => a ∈ X)).
|
||||||
Proof.
|
Proof.
|
||||||
hinduction; try (intros; apply equiv_hprop_allpath ; apply _).
|
hinduction; try (intros; apply equiv_hprop_allpath ; apply _).
|
||||||
- apply (tr (inl idpath)).
|
- apply (tr (inl idpath)).
|
||||||
|
|
|
@ -29,7 +29,7 @@ Section operations_isIn.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Lemma L_isIn (a b : A) :
|
Lemma L_isIn (a b : A) :
|
||||||
isIn a {|b|} -> a = b.
|
a ∈ {|b|} -> a = b.
|
||||||
Proof.
|
Proof.
|
||||||
intros. strip_truncations. assumption.
|
intros. strip_truncations. assumption.
|
||||||
Defined.
|
Defined.
|
||||||
|
@ -63,7 +63,7 @@ Section operations_isIn.
|
||||||
|
|
||||||
(* Union and membership *)
|
(* Union and membership *)
|
||||||
Lemma union_isIn_b (X Y : FSet A) (a : A) :
|
Lemma union_isIn_b (X Y : FSet A) (a : A) :
|
||||||
isIn_b a (U X Y) = orb (isIn_b a X) (isIn_b a Y).
|
isIn_b a (X ∪ Y) = orb (isIn_b a X) (isIn_b a Y).
|
||||||
Proof.
|
Proof.
|
||||||
unfold isIn_b ; unfold dec.
|
unfold isIn_b ; unfold dec.
|
||||||
simpl.
|
simpl.
|
||||||
|
@ -109,7 +109,7 @@ Section SetLattice.
|
||||||
|
|
||||||
Instance fset_max : maximum (FSet A) := U.
|
Instance fset_max : maximum (FSet A) := U.
|
||||||
Instance fset_min : minimum (FSet A) := intersection.
|
Instance fset_min : minimum (FSet A) := intersection.
|
||||||
Instance fset_bot : bottom (FSet A) := E.
|
Instance fset_bot : bottom (FSet A) := ∅.
|
||||||
|
|
||||||
Instance lattice_fset : Lattice (FSet A).
|
Instance lattice_fset : Lattice (FSet A).
|
||||||
Proof.
|
Proof.
|
||||||
|
@ -133,7 +133,7 @@ Section comprehension_properties.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
(** comprehension properties *)
|
(** comprehension properties *)
|
||||||
Lemma comprehension_false Y : comprehension (fun (_ : A) => false) Y = E.
|
Lemma comprehension_false Y : comprehension (fun (_ : A) => false) Y = ∅.
|
||||||
Proof.
|
Proof.
|
||||||
toBool.
|
toBool.
|
||||||
Defined.
|
Defined.
|
||||||
|
@ -145,7 +145,7 @@ Section comprehension_properties.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Lemma comprehension_subset : forall ϕ (X : FSet A),
|
Lemma comprehension_subset : forall ϕ (X : FSet A),
|
||||||
U (comprehension ϕ X) X = X.
|
(comprehension ϕ X) ∪ X = X.
|
||||||
Proof.
|
Proof.
|
||||||
toBool.
|
toBool.
|
||||||
Defined.
|
Defined.
|
||||||
|
@ -159,7 +159,7 @@ Section dec_eq.
|
||||||
Instance fsets_dec_eq : DecidablePaths (FSet A).
|
Instance fsets_dec_eq : DecidablePaths (FSet A).
|
||||||
Proof.
|
Proof.
|
||||||
intros X Y.
|
intros X Y.
|
||||||
apply (decidable_equiv' ((subset Y X) * (subset X Y)) (eq_subset X Y)^-1).
|
apply (decidable_equiv' ((Y ⊆ X) * (X ⊆ Y)) (eq_subset X Y)^-1).
|
||||||
apply decidable_prod.
|
apply decidable_prod.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
|
|
|
@ -34,11 +34,11 @@ Section interface.
|
||||||
|
|
||||||
Class sets :=
|
Class sets :=
|
||||||
{
|
{
|
||||||
f_empty : forall A, f A empty = E ;
|
f_empty : forall A, f A empty = ∅ ;
|
||||||
f_singleton : forall A a, f A (singleton a) = L a;
|
f_singleton : forall A a, f A (singleton a) = {|a|};
|
||||||
f_union : forall A X Y, f A (union X Y) = U (f A X) (f A Y);
|
f_union : forall A X Y, f A (union X Y) = (f A X) ∪ (f A Y);
|
||||||
f_filter : forall A ϕ X, f A (filter ϕ X) = comprehension ϕ (f A X);
|
f_filter : forall A ϕ X, f A (filter ϕ X) = comprehension ϕ (f A X);
|
||||||
f_member : forall A a X, member a X = isIn a (f A X)
|
f_member : forall A a X, member a X = a ∈ (f A X)
|
||||||
}.
|
}.
|
||||||
End interface.
|
End interface.
|
||||||
|
|
||||||
|
@ -47,9 +47,11 @@ Section properties.
|
||||||
Variable (T : Type -> Type) (f : forall A, T A -> FSet A).
|
Variable (T : Type -> Type) (f : forall A, T A -> FSet A).
|
||||||
Context `{sets T f}.
|
Context `{sets T f}.
|
||||||
|
|
||||||
Definition set_eq : forall A, T A -> T A -> hProp := fun A X Y => (BuildhProp (f A X = f A Y)).
|
Definition set_eq : forall A, T A -> T A -> hProp :=
|
||||||
|
fun A X Y => (BuildhProp (f A X = f A Y)).
|
||||||
|
|
||||||
Definition set_subset : forall A, T A -> T A -> hProp := fun A X Y => subset (f A X) (f A Y).
|
Definition set_subset : forall A, T A -> T A -> hProp :=
|
||||||
|
fun A X Y => (f A X) ⊆ (f A Y).
|
||||||
|
|
||||||
Ltac reduce := intros ; repeat (rewrite ?(f_empty _ _) ; rewrite ?(f_singleton _ _) ;
|
Ltac reduce := intros ; repeat (rewrite ?(f_empty _ _) ; rewrite ?(f_singleton _ _) ;
|
||||||
rewrite ?(f_union _ _) ; rewrite ?(f_filter _ _) ;
|
rewrite ?(f_union _ _) ; rewrite ?(f_filter _ _) ;
|
||||||
|
|
|
@ -47,7 +47,7 @@ Section ListToSet.
|
||||||
Context `{Univalence}.
|
Context `{Univalence}.
|
||||||
|
|
||||||
Lemma member_isIn (a : A) (l : list A) :
|
Lemma member_isIn (a : A) (l : list A) :
|
||||||
member a l = isIn a (list_to_set A l).
|
member a l = a ∈ (list_to_set A l).
|
||||||
Proof.
|
Proof.
|
||||||
induction l ; unfold member in * ; simpl in *.
|
induction l ; unfold member in * ; simpl in *.
|
||||||
- reflexivity.
|
- reflexivity.
|
||||||
|
@ -60,7 +60,7 @@ Section ListToSet.
|
||||||
* apply (tr (inr z2)).
|
* apply (tr (inr z2)).
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Definition empty_empty : list_to_set A empty = E := idpath.
|
Definition empty_empty : list_to_set A empty = ∅ := idpath.
|
||||||
|
|
||||||
Lemma filter_comprehension (ϕ : A -> Bool) (l : list A) :
|
Lemma filter_comprehension (ϕ : A -> Bool) (l : list A) :
|
||||||
list_to_set A (filter ϕ l) = comprehension ϕ (list_to_set A l).
|
list_to_set A (filter ϕ l) = comprehension ϕ (list_to_set A l).
|
||||||
|
@ -68,17 +68,17 @@ Section ListToSet.
|
||||||
induction l ; cbn in *.
|
induction l ; cbn in *.
|
||||||
- reflexivity.
|
- reflexivity.
|
||||||
- destruct (ϕ a) ; cbn in * ; unfold list_to_set in IHl.
|
- destruct (ϕ a) ; cbn in * ; unfold list_to_set in IHl.
|
||||||
* refine (ap (fun y => U {|a|} y) _).
|
* refine (ap (fun y => {|a|} ∪ y) _).
|
||||||
apply IHl.
|
apply IHl.
|
||||||
* rewrite nl.
|
* rewrite nl.
|
||||||
apply IHl.
|
apply IHl.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Definition singleton_single (a : A) : list_to_set A (singleton a) = L a :=
|
Definition singleton_single (a : A) : list_to_set A (singleton a) = {|a|} :=
|
||||||
nr (L a).
|
nr {|a|}.
|
||||||
|
|
||||||
Lemma append_union (l1 l2 : list A) :
|
Lemma append_union (l1 l2 : list A) :
|
||||||
list_to_set A (union l1 l2) = U (list_to_set A l1) (list_to_set A l2).
|
list_to_set A (union l1 l2) = (list_to_set A l1) ∪ (list_to_set A l2).
|
||||||
Proof.
|
Proof.
|
||||||
induction l1 ; induction l2 ; cbn.
|
induction l1 ; induction l2 ; cbn.
|
||||||
- apply (union_idem _)^.
|
- apply (union_idem _)^.
|
||||||
|
|
|
@ -62,11 +62,11 @@ Module Export FSet.
|
||||||
comm x y #
|
comm x y #
|
||||||
uP x y px py = uP y x py px).
|
uP x y px py = uP y x py px).
|
||||||
Variable (nlP : forall (x : FSet A) (px: P x),
|
Variable (nlP : forall (x : FSet A) (px: P x),
|
||||||
nl x # uP E x eP px = px).
|
nl x # uP ∅ x eP px = px).
|
||||||
Variable (nrP : forall (x : FSet A) (px: P x),
|
Variable (nrP : forall (x : FSet A) (px: P x),
|
||||||
nr x # uP x E px eP = px).
|
nr x # uP x ∅ px eP = px).
|
||||||
Variable (idemP : forall (x : A),
|
Variable (idemP : forall (x : A),
|
||||||
idem x # uP (L x) (L x) (lP x) (lP x) = lP x).
|
idem x # uP {|x|} {|x|} (lP x) (lP x) = lP x).
|
||||||
|
|
||||||
(* Induction principle *)
|
(* Induction principle *)
|
||||||
Fixpoint FSet_ind
|
Fixpoint FSet_ind
|
||||||
|
|
|
@ -42,29 +42,32 @@ Module Export FSet.
|
||||||
Arguments nl {_} _.
|
Arguments nl {_} _.
|
||||||
Arguments nr {_} _.
|
Arguments nr {_} _.
|
||||||
Arguments idem {_} _.
|
Arguments idem {_} _.
|
||||||
|
Notation "{| x |}" := (L x).
|
||||||
|
Infix "∪" := U (at level 8, right associativity).
|
||||||
|
Notation "∅" := E.
|
||||||
|
|
||||||
Section FSet_induction.
|
Section FSet_induction.
|
||||||
Variable A: Type.
|
Variable A: Type.
|
||||||
Variable (P : FSet A -> Type).
|
Variable (P : FSet A -> Type).
|
||||||
Variable (H : forall a : FSet A, IsHSet (P a)).
|
Variable (H : forall X : FSet A, IsHSet (P X)).
|
||||||
Variable (eP : P E).
|
Variable (eP : P ∅).
|
||||||
Variable (lP : forall a: A, P (L a)).
|
Variable (lP : forall a: A, P {|a|}).
|
||||||
Variable (uP : forall (x y: FSet A), P x -> P y -> P (U x y)).
|
Variable (uP : forall (x y: FSet A), P x -> P y -> P (x ∪ y)).
|
||||||
Variable (assocP : forall (x y z : FSet A)
|
Variable (assocP : forall (x y z : FSet A)
|
||||||
(px: P x) (py: P y) (pz: P z),
|
(px: P x) (py: P y) (pz: P z),
|
||||||
assoc x y z #
|
assoc x y z #
|
||||||
(uP x (U y z) px (uP y z py pz))
|
(uP x (y ∪ z) px (uP y z py pz))
|
||||||
=
|
=
|
||||||
(uP (U x y) z (uP x y px py) pz)).
|
(uP (x ∪ y) z (uP x y px py) pz)).
|
||||||
Variable (commP : forall (x y: FSet A) (px: P x) (py: P y),
|
Variable (commP : forall (x y: FSet A) (px: P x) (py: P y),
|
||||||
comm x y #
|
comm x y #
|
||||||
uP x y px py = uP y x py px).
|
uP x y px py = uP y x py px).
|
||||||
Variable (nlP : forall (x : FSet A) (px: P x),
|
Variable (nlP : forall (x : FSet A) (px: P x),
|
||||||
nl x # uP E x eP px = px).
|
nl x # uP ∅ x eP px = px).
|
||||||
Variable (nrP : forall (x : FSet A) (px: P x),
|
Variable (nrP : forall (x : FSet A) (px: P x),
|
||||||
nr x # uP x E px eP = px).
|
nr x # uP x ∅ px eP = px).
|
||||||
Variable (idemP : forall (x : A),
|
Variable (idemP : forall (x : A),
|
||||||
idem x # uP (L x) (L x) (lP x) (lP x) = lP x).
|
idem x # uP {|x|} {|x|} (lP x) (lP x) = lP x).
|
||||||
|
|
||||||
(* Induction principle *)
|
(* Induction principle *)
|
||||||
Fixpoint FSet_ind
|
Fixpoint FSet_ind
|
||||||
|
@ -183,9 +186,11 @@ Module Export FSet.
|
||||||
|
|
||||||
End FSet_recursion.
|
End FSet_recursion.
|
||||||
|
|
||||||
Instance FSet_recursion A : HitRecursion (FSet A) := {
|
Instance FSet_recursion A : HitRecursion (FSet A) :=
|
||||||
indTy := _; recTy := _;
|
{
|
||||||
H_inductor := FSet_ind A; H_recursor := FSet_rec A }.
|
indTy := _; recTy := _;
|
||||||
|
H_inductor := FSet_ind A; H_recursor := FSet_rec A
|
||||||
|
}.
|
||||||
|
|
||||||
End FSet.
|
End FSet.
|
||||||
|
|
||||||
|
|
Loading…
Reference in New Issue