mirror of
https://github.com/nmvdw/HITs-Examples
synced 2025-11-03 15:13:51 +01:00
Some cleaning in notation
This commit is contained in:
@@ -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 _ _) ;
|
||||
|
||||
@@ -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 _)^.
|
||||
|
||||
Reference in New Issue
Block a user