Some cleaning in notation

This commit is contained in:
Niels 2017-08-07 16:49:46 +02:00
parent 1bab2206a3
commit 1e373364b2
9 changed files with 136 additions and 137 deletions

View File

@ -8,7 +8,7 @@ Section ext.
Context `{Univalence}.
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.
intros X Y.
eapply equiv_iff_hprop_uncurried.
@ -20,8 +20,8 @@ Section ext.
Defined.
Lemma subset_isIn (X Y : FSet A) :
(forall (a : A), isIn a X -> isIn a Y)
<~> (subset X Y).
(forall (a : A), a X -> a Y)
<~> (X Y).
Proof.
eapply equiv_iff_hprop_uncurried.
split.
@ -32,19 +32,17 @@ Section ext.
apply f.
apply tr ; reflexivity.
+ intros X1 X2 H1 H2 f.
enough (subset X1 Y).
enough (subset X2 Y).
enough (X1 Y).
enough (X2 Y).
{ split. apply X. apply X0. }
* apply H2.
intros a Ha.
apply f.
apply tr.
refine (f _ (tr _)).
right.
apply Ha.
* apply H1.
intros a Ha.
apply f.
apply tr.
refine (f _ (tr _)).
left.
apply Ha.
- hinduction X ;
@ -64,7 +62,7 @@ Section ext.
(** ** 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.
unshelve eapply BuildEquiv.
{ intro H'. rewrite H'. split; apply union_idem. }
@ -76,20 +74,20 @@ Section ext.
Defined.
Lemma eq_subset (X Y : FSet A) :
X = Y <~> (subset Y X * subset X Y).
X = Y <~> (Y X * X Y).
Proof.
transitivity ((U Y X = X) * (U X Y = Y)).
transitivity ((Y X = X) * (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).
X = Y <~> (forall (a : A), a X = 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)) _ _ _).
refine (@equiv_compose' _ ((forall a, a Y -> a X)
*(forall a, a X -> a Y)) _ _ _).
- apply equiv_iff_hprop_uncurried.
split.
* intros [H1 H2 a].

View File

@ -11,7 +11,7 @@ Section Length.
Opaque isIn_b.
Definition length (x: FSetC A) : nat.
Definition length (x : FSetC A) : nat.
Proof.
simple refine (FSetC_ind A _ _ _ _ _ _ x ); simpl.
- exact 0.
@ -31,7 +31,7 @@ Section Length.
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.
intro a.
cbn. reflexivity.

View File

@ -3,84 +3,78 @@ Require Import HoTT HitTactics.
Require Import representations.definition disjunction lattice.
Section operations.
Context {A : Type}.
Context `{Univalence}.
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 ; symmetry ; apply lor_assoc.
- apply lor_commutative.
- apply lor_nl.
- apply lor_nr.
- intros ; apply lor_idem.
Defined.
Definition isIn : A -> FSet A -> hProp.
Proof.
intros a.
hrecursion.
- exists Empty.
exact _.
- intro a'.
exists (Trunc (-1) (a = a')).
exact _.
- apply lor.
- intros ; symmetry ; apply lor_assoc.
- apply lor_commutative.
- apply lor_nl.
- apply lor_nr.
- intros ; apply lor_idem.
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 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 comprehension :
(A -> Bool) -> FSet A -> FSet A.
Proof.
intros P.
hrecursion.
- apply .
- intro a.
refine (if (P a) then {|a|} else ).
- 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 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.
Definition isEmpty :
FSet A -> Bool.
Proof.
simple refine (FSet_rec _ _ _ true (fun _ => false) andb _ _ _ _ _)
; try eauto with bool_lattice_hints typeclass_instances.
intros ; symmetry ; eauto with lattice_hints typeclass_instances.
Defined.
End operations.

View File

@ -7,7 +7,7 @@ Section operations_isIn.
Context {A : Type}.
Context `{Univalence}.
Lemma union_idem : forall x: FSet A, U x x = x.
Lemma union_idem : forall x: FSet A, x x = x.
Proof.
hinduction ; try (intros ; apply set_path2).
- apply nl.
@ -24,7 +24,7 @@ Section operations_isIn.
(** ** Properties about subset relation. *)
Lemma subset_union (X Y : FSet A) :
subset X Y -> U X Y = Y.
X Y -> X Y = Y.
Proof.
hinduction X ; try (intros; apply path_forall; intro; apply set_path2).
- intros. apply nl.
@ -54,7 +54,7 @@ Section operations_isIn.
Defined.
Lemma subset_union_l (X : FSet A) :
forall Y, subset X (U X Y).
forall Y, X X Y.
Proof.
hinduction X ; try (repeat (intro; intros; apply path_forall);
intro ; apply equiv_hprop_allpath ; apply _).
@ -69,8 +69,8 @@ Section operations_isIn.
(* simplify it using extensionality *)
Lemma comprehension_or : forall ϕ ψ (x: FSet A),
comprehension (fun a => orb (ϕ a) (ψ a)) x = U (comprehension ϕ x)
(comprehension ψ x).
comprehension (fun a => orb (ϕ a) (ψ a)) x
= (comprehension ϕ x) (comprehension ψ x).
Proof.
intros ϕ ψ.
hinduction ; try (intros; apply set_path2).
@ -101,15 +101,15 @@ Section properties.
Context `{Univalence}.
(** 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)
: 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,
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.
hinduction ; try (intros ; apply set_path2) ; cbn.
- destruct (ϕ a) ; reflexivity.
@ -139,7 +139,7 @@ Section properties.
(* The proof can be simplified using extensionality *)
(** comprehension properties *)
Lemma comprehension_false Y : comprehension (fun (_ : A) => false) Y = E.
Lemma comprehension_false Y : comprehension (fun (_ : A) => false) Y = .
Proof.
hrecursion Y; try (intros; apply set_path2).
- reflexivity.
@ -151,7 +151,7 @@ Section properties.
(* Can be simplified using extensionality *)
Lemma comprehension_subset : forall ϕ (X : FSet A),
U (comprehension ϕ X) X = X.
(comprehension ϕ X) X = X.
Proof.
intros ϕ.
hrecursion; try (intros ; apply set_path2) ; cbn.
@ -171,7 +171,7 @@ Section properties.
reflexivity.
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.
hinduction; try (intros; apply equiv_hprop_allpath ; apply _).
- apply (tr (inl idpath)).

View File

@ -29,7 +29,7 @@ Section operations_isIn.
Defined.
Lemma L_isIn (a b : A) :
isIn a {|b|} -> a = b.
a {|b|} -> a = b.
Proof.
intros. strip_truncations. assumption.
Defined.
@ -63,7 +63,7 @@ Section operations_isIn.
(* Union and membership *)
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.
unfold isIn_b ; unfold dec.
simpl.
@ -109,7 +109,7 @@ Section SetLattice.
Instance fset_max : maximum (FSet A) := U.
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).
Proof.
@ -133,7 +133,7 @@ Section comprehension_properties.
Defined.
(** comprehension properties *)
Lemma comprehension_false Y : comprehension (fun (_ : A) => false) Y = E.
Lemma comprehension_false Y : comprehension (fun (_ : A) => false) Y = .
Proof.
toBool.
Defined.
@ -145,7 +145,7 @@ Section comprehension_properties.
Defined.
Lemma comprehension_subset : forall ϕ (X : FSet A),
U (comprehension ϕ X) X = X.
(comprehension ϕ X) X = X.
Proof.
toBool.
Defined.
@ -159,7 +159,7 @@ Section dec_eq.
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 (decidable_equiv' ((Y X) * (X Y)) (eq_subset X Y)^-1).
apply decidable_prod.
Defined.

View File

@ -34,11 +34,11 @@ Section interface.
Class sets :=
{
f_empty : forall A, f A empty = E ;
f_singleton : forall A a, f A (singleton a) = L a;
f_union : forall A X Y, f A (union X Y) = U (f A X) (f A Y);
f_empty : forall A, f A empty = ;
f_singleton : forall A a, f A (singleton a) = {|a|};
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_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.
@ -47,9 +47,11 @@ Section properties.
Variable (T : Type -> Type) (f : forall A, T A -> FSet A).
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 _ _) ;
rewrite ?(f_union _ _) ; rewrite ?(f_filter _ _) ;

View File

@ -47,7 +47,7 @@ Section ListToSet.
Context `{Univalence}.
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.
induction l ; unfold member in * ; simpl in *.
- reflexivity.
@ -60,7 +60,7 @@ Section ListToSet.
* apply (tr (inr z2)).
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) :
list_to_set A (filter ϕ l) = comprehension ϕ (list_to_set A l).
@ -68,17 +68,17 @@ Section ListToSet.
induction l ; cbn in *.
- reflexivity.
- 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.
* rewrite nl.
apply IHl.
Defined.
Definition singleton_single (a : A) : list_to_set A (singleton a) = L a :=
nr (L a).
Definition singleton_single (a : A) : list_to_set A (singleton a) = {|a|} :=
nr {|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.
induction l1 ; induction l2 ; cbn.
- apply (union_idem _)^.

View File

@ -62,11 +62,11 @@ Module Export FSet.
comm x y #
uP x y px py = uP y x py px).
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),
nr x # uP x E px eP = px).
nr x # uP x px eP = px).
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 *)
Fixpoint FSet_ind

View File

@ -42,29 +42,32 @@ Module Export FSet.
Arguments nl {_} _.
Arguments nr {_} _.
Arguments idem {_} _.
Notation "{| x |}" := (L x).
Infix "" := U (at level 8, right associativity).
Notation "" := E.
Section FSet_induction.
Variable A: Type.
Variable (P : FSet A -> Type).
Variable (H : forall a : FSet A, IsHSet (P a)).
Variable (eP : P E).
Variable (lP : forall a: A, P (L a)).
Variable (uP : forall (x y: FSet A), P x -> P y -> P (U x y)).
Variable (H : forall X : FSet A, IsHSet (P X)).
Variable (eP : P ).
Variable (lP : forall a: A, P {|a|}).
Variable (uP : forall (x y: FSet A), P x -> P y -> P (x y)).
Variable (assocP : forall (x y z : FSet A)
(px: P x) (py: P y) (pz: P 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),
comm x y #
uP x y px py = uP y x py px).
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),
nr x # uP x E px eP = px).
nr x # uP x px eP = px).
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 *)
Fixpoint FSet_ind
@ -183,9 +186,11 @@ Module Export FSet.
End FSet_recursion.
Instance FSet_recursion A : HitRecursion (FSet A) := {
indTy := _; recTy := _;
H_inductor := FSet_ind A; H_recursor := FSet_rec A }.
Instance FSet_recursion A : HitRecursion (FSet A) :=
{
indTy := _; recTy := _;
H_inductor := FSet_ind A; H_recursor := FSet_rec A
}.
End FSet.