2017-09-07 15:19:48 +02:00
|
|
|
|
(** Classes for set operations, so they can be overloaded. *)
|
2017-08-08 15:30:12 +02:00
|
|
|
|
Require Import HoTT.
|
|
|
|
|
|
2017-09-21 14:12:51 +02:00
|
|
|
|
(** The operations on sets for which we add names. *)
|
2017-08-08 15:30:12 +02:00
|
|
|
|
Section structure.
|
2017-08-08 17:00:30 +02:00
|
|
|
|
Variable (T A : Type).
|
2017-08-08 15:30:12 +02:00
|
|
|
|
|
|
|
|
|
Class hasMembership : Type :=
|
2017-08-08 17:00:30 +02:00
|
|
|
|
member : A -> T -> hProp.
|
|
|
|
|
|
|
|
|
|
Class hasMembership_decidable : Type :=
|
|
|
|
|
member_dec : A -> T -> Bool.
|
|
|
|
|
|
|
|
|
|
Class hasSubset : Type :=
|
|
|
|
|
subset : T -> T -> hProp.
|
|
|
|
|
|
|
|
|
|
Class hasSubset_decidable : Type :=
|
|
|
|
|
subset_dec : T -> T -> Bool.
|
2017-08-08 15:30:12 +02:00
|
|
|
|
|
|
|
|
|
Class hasEmpty : Type :=
|
2017-08-08 17:00:30 +02:00
|
|
|
|
empty : T.
|
2017-08-08 15:30:12 +02:00
|
|
|
|
|
|
|
|
|
Class hasSingleton : Type :=
|
2017-08-08 17:00:30 +02:00
|
|
|
|
singleton : A -> T.
|
2017-08-08 15:30:12 +02:00
|
|
|
|
|
|
|
|
|
Class hasUnion : Type :=
|
2017-08-08 17:00:30 +02:00
|
|
|
|
union : T -> T -> T.
|
2017-08-08 15:30:12 +02:00
|
|
|
|
|
|
|
|
|
Class hasIntersection : Type :=
|
2017-08-08 17:00:30 +02:00
|
|
|
|
intersection : T -> T -> T.
|
2017-08-08 15:30:12 +02:00
|
|
|
|
|
|
|
|
|
Class hasComprehension : Type :=
|
2017-08-08 17:00:30 +02:00
|
|
|
|
filter : (A -> Bool) -> T -> T.
|
2017-08-08 15:30:12 +02:00
|
|
|
|
End structure.
|
|
|
|
|
|
|
|
|
|
Arguments member {_} {_} {_} _ _.
|
2017-08-08 17:00:30 +02:00
|
|
|
|
Arguments subset {_} {_} _ _.
|
|
|
|
|
Arguments member_dec {_} {_} {_} _ _.
|
|
|
|
|
Arguments subset_dec {_} {_} _ _.
|
|
|
|
|
Arguments empty {_} {_}.
|
2017-08-08 15:30:12 +02:00
|
|
|
|
Arguments singleton {_} {_} {_} _.
|
2017-08-08 17:00:30 +02:00
|
|
|
|
Arguments union {_} {_} _ _.
|
|
|
|
|
Arguments intersection {_} {_} _ _.
|
2017-08-08 15:30:12 +02:00
|
|
|
|
Arguments filter {_} {_} {_} _ _.
|
|
|
|
|
|
|
|
|
|
Notation "∅" := empty.
|
|
|
|
|
Notation "{| x |}" := (singleton x).
|
|
|
|
|
Infix "∪" := union (at level 8, right associativity).
|
|
|
|
|
Notation "(∪)" := union (only parsing).
|
|
|
|
|
Notation "( X ∪)" := (union X) (only parsing).
|
|
|
|
|
Notation "( ∪ Y )" := (fun X => X ∪ Y) (only parsing).
|
|
|
|
|
Infix "∩" := intersection (at level 8, right associativity).
|
|
|
|
|
Notation "( ∩ )" := intersection (only parsing).
|
|
|
|
|
Notation "( X ∩ )" := (intersection X) (only parsing).
|
|
|
|
|
Notation "( ∩ Y )" := (fun X => X ∩ Y) (only parsing).
|
2017-11-01 17:47:41 +01:00
|
|
|
|
Notation "{| X | ϕ |}" := (filter ϕ X).
|
2017-08-08 17:00:30 +02:00
|
|
|
|
Infix "∈" := member (at level 9, right associativity).
|
|
|
|
|
Infix "⊆" := subset (at level 10, right associativity).
|
|
|
|
|
Infix "∈_d" := member_dec (at level 9, right associativity).
|
2017-11-01 17:47:41 +01:00
|
|
|
|
Infix "⊆_d" := subset_dec (at level 10, right associativity).
|