Skip to content
Snippets Groups Projects
Commit ad870687 authored by Robbert Krebbers's avatar Robbert Krebbers
Browse files

Type class for ⊤ to get overloaded notation for entire set.

parent 882fdc9d
No related branches found
No related tags found
No related merge requests found
......@@ -209,6 +209,9 @@ intersection [(∩)], and difference [(∖)], the singleton [{[_]}], the subset
Class Empty A := empty: A.
Notation "∅" := empty : C_scope.
Class Top A := top : A.
Notation "⊤" := top : C_scope.
Class Union A := union: A A A.
Instance: Params (@union) 2.
Infix "∪" := union (at level 50, left associativity) : C_scope.
......@@ -311,7 +314,7 @@ Instance: Params (@disjoint) 2.
Infix "⊥" := disjoint (at level 70) : C_scope.
Notation "(⊥)" := disjoint (only parsing) : C_scope.
Notation "( X ⊥.)" := (disjoint X) (only parsing) : C_scope.
Notation "(.⊥ X )" := (λ Y, Y X) (only parsing) : C_scope.
Notation "(.⊥ X )" := (λ Y, Y X) (only parsing) : C_scope.
Infix "⊥*" := (Forall2 ()) (at level 70) : C_scope.
Notation "(⊥*)" := (Forall2 ()) (only parsing) : C_scope.
Infix "⊥**" := (Forall2 (⊥*)) (at level 70) : C_scope.
......
......@@ -6,7 +6,7 @@ From stdpp Require Export prelude.
Record bset (A : Type) : Type := mkBSet { bset_car : A bool }.
Arguments mkBSet {_} _.
Arguments bset_car {_} _ _.
Definition bset_all {A} : bset A := mkBSet (λ _, true).
Instance bset_top {A} : Top (bset A) := mkBSet (λ _, true).
Instance bset_empty {A} : Empty (bset A) := mkBSet (λ _, false).
Instance bset_singleton {A} `{ x y : A, Decision (x = y)} :
Singleton A (bset A) := λ x, mkBSet (λ y, bool_decide (y = x)).
......
......@@ -148,7 +148,7 @@ Instance coPset_singleton : Singleton positive coPset := λ p,
coPset_singleton_raw p coPset_singleton_wf _.
Instance coPset_elem_of : ElemOf positive coPset := λ p X, e_of p (`X).
Instance coPset_empty : Empty coPset := coPLeaf false I.
Definition coPset_all : coPset := coPLeaf true I.
Instance coPset_top : Top coPset := coPLeaf true I.
Instance coPset_union : Union coPset := λ X Y,
let (t1,Ht1) := X in let (t2,Ht2) := Y in
(t1 t2) coPset_union_wf _ _ Ht1 Ht2.
......
......@@ -6,7 +6,7 @@ From stdpp Require Export prelude.
Record set (A : Type) : Type := mkSet { set_car : A Prop }.
Arguments mkSet {_} _.
Arguments set_car {_} _ _.
Definition set_all {A} : set A := mkSet (λ _, True).
Instance set_all {A} : Top (set A) := mkSet (λ _, True).
Instance set_empty {A} : Empty (set A) := mkSet (λ _, False).
Instance set_singleton {A} : Singleton A (set A) := λ x, mkSet (x =).
Instance set_elem_of {A} : ElemOf A (set A) := λ x X, set_car X x.
......
0% Loading or .
You are about to add 0 people to the discussion. Proceed with caution.
Finish editing this message first!
Please register or to comment