mirror of https://github.com/nmvdw/HITs-Examples
Completely fixed notation
This commit is contained in:
parent
92bc50d79f
commit
3cda0d9bf2
|
@ -45,7 +45,7 @@ Section ext.
|
||||||
- intros a Y.
|
- intros a Y.
|
||||||
apply (tr(inl(tr idpath))).
|
apply (tr(inl(tr idpath))).
|
||||||
- intros X1 X2 HX1 HX2 Y.
|
- intros X1 X2 HX1 HX2 Y.
|
||||||
split.
|
split ; unfold subset in *.
|
||||||
* rewrite <- assoc. apply HX1.
|
* rewrite <- assoc. apply HX1.
|
||||||
* rewrite (comm X1 X2). rewrite <- assoc. apply HX2.
|
* rewrite (comm X1 X2). rewrite <- assoc. apply HX2.
|
||||||
Defined.
|
Defined.
|
||||||
|
@ -64,7 +64,7 @@ Section ext.
|
||||||
|
|
||||||
Lemma subset_isIn (X Y : FSet A) :
|
Lemma subset_isIn (X Y : FSet A) :
|
||||||
(forall (a : A), a ∈ X -> a ∈ Y)
|
(forall (a : A), a ∈ X -> a ∈ Y)
|
||||||
<~> (X ⊆ Y).
|
<~> X ⊆ Y.
|
||||||
Proof.
|
Proof.
|
||||||
eapply equiv_iff_hprop_uncurried.
|
eapply equiv_iff_hprop_uncurried.
|
||||||
split.
|
split.
|
||||||
|
@ -149,4 +149,4 @@ Section ext.
|
||||||
split ; apply subset_isIn.
|
split ; apply subset_isIn.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
End ext.
|
End ext.
|
|
@ -45,7 +45,7 @@ Section Iso.
|
||||||
intros x.
|
intros x.
|
||||||
hrecursion x; try (intros; apply path_forall; intro; apply set_path2).
|
hrecursion x; try (intros; apply path_forall; intro; apply set_path2).
|
||||||
- intros. symmetry. apply nl.
|
- intros. symmetry. apply nl.
|
||||||
- intros a x HR y. rewrite HR. apply assoc.
|
- intros a x HR y. unfold union, fsetc_union in *. rewrite HR. apply assoc.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Lemma repr_iso_id_l: forall (x: FSet A), FSetC_to_FSet (FSet_to_FSetC x) = x.
|
Lemma repr_iso_id_l: forall (x: FSet A), FSetC_to_FSet (FSet_to_FSetC x) = x.
|
||||||
|
@ -56,7 +56,6 @@ Section Iso.
|
||||||
- intros x y p q. rewrite append_union, p, q. reflexivity.
|
- intros x y p q. rewrite append_union, p, q. reflexivity.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
|
|
||||||
Lemma repr_iso_id_r: forall (x: FSetC A), FSet_to_FSetC (FSetC_to_FSet x) = x.
|
Lemma repr_iso_id_r: forall (x: FSetC A), FSet_to_FSetC (FSetC_to_FSet x) = x.
|
||||||
Proof.
|
Proof.
|
||||||
hinduction; try (intros; apply set_path2).
|
hinduction; try (intros; apply set_path2).
|
||||||
|
|
|
@ -9,24 +9,22 @@ Section Length.
|
||||||
Context {A_deceq : DecidablePaths A}.
|
Context {A_deceq : DecidablePaths A}.
|
||||||
Context `{Univalence}.
|
Context `{Univalence}.
|
||||||
|
|
||||||
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.
|
||||||
- intros a y n.
|
- intros a y n.
|
||||||
pose (y' := FSetC_to_FSet y).
|
pose (y' := FSetC_to_FSet y).
|
||||||
exact (if isIn_b a y' then n else (S n)).
|
exact (if a ∈_d y' then n else (S n)).
|
||||||
- intros. rewrite transport_const. cbn.
|
- intros. rewrite transport_const. simpl.
|
||||||
simplify_isIn_b. simpl. reflexivity.
|
simplify_isIn_b. reflexivity.
|
||||||
- intros. rewrite transport_const. cbn.
|
- intros. rewrite transport_const. simpl.
|
||||||
simplify_isIn_b.
|
simplify_isIn_b.
|
||||||
destruct (dec (a = b)) as [Hab | Hab].
|
destruct (dec (a = b)) as [Hab | Hab].
|
||||||
+ rewrite Hab. simplify_isIn_b. simpl. reflexivity.
|
+ rewrite Hab. simplify_isIn_b. reflexivity.
|
||||||
+ rewrite ?L_isIn_b_false; auto.
|
+ rewrite ?L_isIn_b_false; auto.
|
||||||
++ simpl.
|
++ simpl.
|
||||||
destruct (isIn_b a (FSetC_to_FSet x0)), (isIn_b b (FSetC_to_FSet x0))
|
destruct (a ∈_d (FSetC_to_FSet x0)), (b ∈_d (FSetC_to_FSet x0))
|
||||||
; reflexivity.
|
; reflexivity.
|
||||||
++ intro p. contradiction (Hab p^).
|
++ intro p. contradiction (Hab p^).
|
||||||
Defined.
|
Defined.
|
||||||
|
|
|
@ -3,12 +3,11 @@ Require Import HoTT HitTactics.
|
||||||
Require Import representations.definition disjunction lattice.
|
Require Import representations.definition disjunction lattice.
|
||||||
|
|
||||||
Section operations.
|
Section operations.
|
||||||
Context {A : Type}.
|
|
||||||
Context `{Univalence}.
|
Context `{Univalence}.
|
||||||
|
|
||||||
Definition isIn : A -> FSet A -> hProp.
|
Global Instance fset_member : forall A, hasMembership (FSet A) A.
|
||||||
Proof.
|
Proof.
|
||||||
intros a.
|
intros A a.
|
||||||
hrecursion.
|
hrecursion.
|
||||||
- exists Empty.
|
- exists Empty.
|
||||||
exact _.
|
exact _.
|
||||||
|
@ -23,10 +22,9 @@ Section operations.
|
||||||
- intros ; apply lor_idem.
|
- intros ; apply lor_idem.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Definition comprehension :
|
Global Instance fset_comprehension : forall A, hasComprehension (FSet A) A.
|
||||||
(A -> Bool) -> FSet A -> FSet A.
|
|
||||||
Proof.
|
Proof.
|
||||||
intros P.
|
intros A P.
|
||||||
hrecursion.
|
hrecursion.
|
||||||
- apply ∅.
|
- apply ∅.
|
||||||
- intro a.
|
- intro a.
|
||||||
|
@ -42,7 +40,7 @@ Section operations.
|
||||||
+ apply nl.
|
+ apply nl.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Definition isEmpty :
|
Definition isEmpty (A : Type) :
|
||||||
FSet A -> Bool.
|
FSet A -> Bool.
|
||||||
Proof.
|
Proof.
|
||||||
simple refine (FSet_rec _ _ _ true (fun _ => false) andb _ _ _ _ _)
|
simple refine (FSet_rec _ _ _ true (fun _ => false) andb _ _ _ _ _)
|
||||||
|
@ -50,7 +48,7 @@ Section operations.
|
||||||
intros ; symmetry ; eauto with lattice_hints typeclass_instances.
|
intros ; symmetry ; eauto with lattice_hints typeclass_instances.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Definition single_product {B : Type} (a : A) : FSet B -> FSet (A * B).
|
Definition single_product {A B : Type} (a : A) : FSet B -> FSet (A * B).
|
||||||
Proof.
|
Proof.
|
||||||
hrecursion.
|
hrecursion.
|
||||||
- apply ∅.
|
- apply ∅.
|
||||||
|
@ -64,7 +62,7 @@ Section operations.
|
||||||
- intros ; apply idem.
|
- intros ; apply idem.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Definition product {B : Type} : FSet A -> FSet B -> FSet (A * B).
|
Definition product {A B : Type} : FSet A -> FSet B -> FSet (A * B).
|
||||||
Proof.
|
Proof.
|
||||||
intros X Y.
|
intros X Y.
|
||||||
hrecursion X.
|
hrecursion X.
|
||||||
|
@ -78,29 +76,10 @@ Section operations.
|
||||||
- intros ; apply nr.
|
- intros ; apply nr.
|
||||||
- intros ; apply union_idem.
|
- intros ; apply union_idem.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
End operations.
|
Global Instance fset_subset : forall A, hasSubset (FSet A).
|
||||||
|
|
||||||
Section instances_operations.
|
|
||||||
Global Instance fset_comprehension : hasComprehension FSet.
|
|
||||||
Proof.
|
Proof.
|
||||||
intros A ϕ X.
|
intros A X Y.
|
||||||
apply (comprehension ϕ X).
|
|
||||||
Defined.
|
|
||||||
|
|
||||||
Context `{Univalence}.
|
|
||||||
|
|
||||||
Global Instance fset_member : hasMembership FSet.
|
|
||||||
Proof.
|
|
||||||
intros A a X.
|
|
||||||
apply (isIn a X).
|
|
||||||
Defined.
|
|
||||||
|
|
||||||
Context {A : Type}.
|
|
||||||
|
|
||||||
Definition subset : FSet A -> FSet A -> hProp.
|
|
||||||
Proof.
|
|
||||||
intros X Y.
|
|
||||||
hrecursion X.
|
hrecursion X.
|
||||||
- exists Unit.
|
- exists Unit.
|
||||||
exact _.
|
exact _.
|
||||||
|
@ -123,6 +102,4 @@ Section instances_operations.
|
||||||
* intros p.
|
* intros p.
|
||||||
split ; apply p.
|
split ; apply p.
|
||||||
Defined.
|
Defined.
|
||||||
End instances_operations.
|
End operations.
|
||||||
|
|
||||||
Infix "⊆" := subset (at level 10, right associativity).
|
|
|
@ -3,8 +3,7 @@ Require Import HoTT HitTactics.
|
||||||
Require Import representations.cons_repr.
|
Require Import representations.cons_repr.
|
||||||
|
|
||||||
Section operations.
|
Section operations.
|
||||||
|
Global Instance fsetc_union : forall A, hasUnion (FSetC A).
|
||||||
Global Instance fsetc_union : hasUnion FSetC.
|
|
||||||
Proof.
|
Proof.
|
||||||
intros A x y.
|
intros A x y.
|
||||||
hinduction x.
|
hinduction x.
|
||||||
|
@ -14,6 +13,6 @@ Section operations.
|
||||||
- apply comm.
|
- apply comm.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Global Instance fsetc_singleton : hasSingleton FSetC := fun A a => a;;∅.
|
Global Instance fsetc_singleton : forall A, hasSingleton (FSetC A) A := fun A a => a;;∅.
|
||||||
|
|
||||||
End operations.
|
End operations.
|
|
@ -6,8 +6,9 @@ Section decidable_A.
|
||||||
Context {A : Type}.
|
Context {A : Type}.
|
||||||
Context {A_deceq : DecidablePaths A}.
|
Context {A_deceq : DecidablePaths A}.
|
||||||
Context `{Univalence}.
|
Context `{Univalence}.
|
||||||
|
|
||||||
Global Instance isIn_decidable : forall (a : A) (X : FSet A), Decidable (a ∈ X).
|
Global Instance isIn_decidable : forall (a : A) (X : FSet A),
|
||||||
|
Decidable (a ∈ X).
|
||||||
Proof.
|
Proof.
|
||||||
intros a.
|
intros a.
|
||||||
hinduction ; try (intros ; apply path_ishprop).
|
hinduction ; try (intros ; apply path_ishprop).
|
||||||
|
@ -23,26 +24,31 @@ Section decidable_A.
|
||||||
- intros. apply _.
|
- intros. apply _.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Definition isIn_b (a : A) (X : FSet A) :=
|
Global Instance fset_member_bool : hasMembership_decidable (FSet A) A.
|
||||||
match dec (a ∈ X) with
|
Proof.
|
||||||
| inl _ => true
|
intros a X.
|
||||||
| inr _ => false
|
destruct (dec (a ∈ X)).
|
||||||
end.
|
- apply true.
|
||||||
|
- apply false.
|
||||||
|
Defined.
|
||||||
|
|
||||||
Global Instance subset_decidable : forall (X Y : FSet A), Decidable (X ⊆ Y).
|
Global Instance subset_decidable : forall (X Y : FSet A), Decidable (X ⊆ Y).
|
||||||
Proof.
|
Proof.
|
||||||
hinduction ; try (intros ; apply path_ishprop).
|
hinduction ; try (intros ; apply path_ishprop).
|
||||||
- intro ; apply _.
|
- intro ; apply _.
|
||||||
- intros. apply _.
|
- intros ; apply _.
|
||||||
- intros ; apply _.
|
- intros. unfold subset in *. apply _.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Definition subset_b (X Y : FSet A) :=
|
Global Instance fset_subset_bool : hasSubset_decidable (FSet A).
|
||||||
match dec (X ⊆ Y) with
|
Proof.
|
||||||
| inl _ => true
|
intros X Y.
|
||||||
| inr _ => false
|
destruct (dec (X ⊆ Y)).
|
||||||
end.
|
- apply true.
|
||||||
|
- apply false.
|
||||||
|
Defined.
|
||||||
|
|
||||||
Definition intersection (X Y : FSet A) : FSet A := comprehension (fun a => isIn_b a Y) X.
|
Global Instance fset_intersection : hasIntersection (FSet A)
|
||||||
|
:= fun X Y => {|X & (fun a => a ∈_d Y)|}.
|
||||||
|
|
||||||
End decidable_A.
|
End decidable_A.
|
|
@ -150,7 +150,7 @@ Section properties.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
(** comprehension properties *)
|
(** comprehension properties *)
|
||||||
Lemma comprehension_false X : {|X & (fun (_ : A) => false)|} = ∅.
|
Lemma comprehension_false : forall (X : FSet A), {|X & fun _ => false|} = ∅.
|
||||||
Proof.
|
Proof.
|
||||||
toHProp.
|
toHProp.
|
||||||
Defined.
|
Defined.
|
||||||
|
@ -172,7 +172,7 @@ Section properties.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Lemma comprehension_all : forall (X : FSet A),
|
Lemma comprehension_all : forall (X : FSet A),
|
||||||
comprehension (fun a => true) X = X.
|
{|X & fun _ => true|} = X.
|
||||||
Proof.
|
Proof.
|
||||||
toHProp.
|
toHProp.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
|
@ -8,13 +8,13 @@ Section operations_isIn.
|
||||||
|
|
||||||
Context {A : Type} `{DecidablePaths A} `{Univalence}.
|
Context {A : Type} `{DecidablePaths A} `{Univalence}.
|
||||||
|
|
||||||
Lemma ext : forall (S T : FSet A), (forall a, isIn_b a S = isIn_b a T) -> S = T.
|
Lemma ext : forall (S T : FSet A), (forall a, a ∈_d S = a ∈_d T) -> S = T.
|
||||||
Proof.
|
Proof.
|
||||||
intros X Y H2.
|
intros X Y H2.
|
||||||
apply fset_ext.
|
apply fset_ext.
|
||||||
intro a.
|
intro a.
|
||||||
specialize (H2 a).
|
specialize (H2 a).
|
||||||
unfold isIn_b, dec in H2.
|
unfold member_dec, fset_member_bool, dec in H2.
|
||||||
destruct (isIn_decidable a X) ; destruct (isIn_decidable a Y).
|
destruct (isIn_decidable a X) ; destruct (isIn_decidable a Y).
|
||||||
- apply path_iff_hprop ; intro ; assumption.
|
- apply path_iff_hprop ; intro ; assumption.
|
||||||
- contradiction (true_ne_false).
|
- contradiction (true_ne_false).
|
||||||
|
@ -23,7 +23,7 @@ Section operations_isIn.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Lemma empty_isIn (a : A) :
|
Lemma empty_isIn (a : A) :
|
||||||
isIn_b a ∅ = false.
|
a ∈_d ∅ = false.
|
||||||
Proof.
|
Proof.
|
||||||
reflexivity.
|
reflexivity.
|
||||||
Defined.
|
Defined.
|
||||||
|
@ -35,9 +35,9 @@ Section operations_isIn.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Lemma L_isIn_b_true (a b : A) (p : a = b) :
|
Lemma L_isIn_b_true (a b : A) (p : a = b) :
|
||||||
isIn_b a {|b|} = true.
|
a ∈_d {|b|} = true.
|
||||||
Proof.
|
Proof.
|
||||||
unfold isIn_b, dec.
|
unfold member_dec, fset_member_bool, dec.
|
||||||
destruct (isIn_decidable a {|b|}) as [n | n] .
|
destruct (isIn_decidable a {|b|}) as [n | n] .
|
||||||
- reflexivity.
|
- reflexivity.
|
||||||
- simpl in n.
|
- simpl in n.
|
||||||
|
@ -45,15 +45,15 @@ Section operations_isIn.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Lemma L_isIn_b_aa (a : A) :
|
Lemma L_isIn_b_aa (a : A) :
|
||||||
isIn_b a {|a|} = true.
|
a ∈_d {|a|} = true.
|
||||||
Proof.
|
Proof.
|
||||||
apply L_isIn_b_true ; reflexivity.
|
apply L_isIn_b_true ; reflexivity.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Lemma L_isIn_b_false (a b : A) (p : a <> b) :
|
Lemma L_isIn_b_false (a b : A) (p : a <> b) :
|
||||||
isIn_b a {|b|} = false.
|
a ∈_d {|b|} = false.
|
||||||
Proof.
|
Proof.
|
||||||
unfold isIn_b, dec.
|
unfold member_dec, fset_member_bool, dec in *.
|
||||||
destruct (isIn_decidable a {|b|}).
|
destruct (isIn_decidable a {|b|}).
|
||||||
- simpl in t.
|
- simpl in t.
|
||||||
strip_truncations.
|
strip_truncations.
|
||||||
|
@ -63,24 +63,25 @@ 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 (X ∪ Y) = orb (isIn_b a X) (isIn_b a Y).
|
a ∈_d (X ∪ Y) = orb (a ∈_d X) (a ∈_d Y).
|
||||||
Proof.
|
Proof.
|
||||||
unfold isIn_b ; unfold dec.
|
unfold member_dec, fset_member_bool, dec.
|
||||||
simpl.
|
simpl.
|
||||||
destruct (isIn_decidable a X) ; destruct (isIn_decidable a Y) ; reflexivity.
|
destruct (isIn_decidable a X) ; destruct (isIn_decidable a Y) ; reflexivity.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Lemma comprehension_isIn_b (Y : FSet A) (ϕ : A -> Bool) (a : A) :
|
Lemma comprehension_isIn_b (Y : FSet A) (ϕ : A -> Bool) (a : A) :
|
||||||
isIn_b a {|Y & ϕ|} = andb (isIn_b a Y) (ϕ a).
|
a ∈_d {|Y & ϕ|} = andb (a ∈_d Y) (ϕ a).
|
||||||
Proof.
|
Proof.
|
||||||
unfold isIn_b, dec ; simpl.
|
unfold member_dec, fset_member_bool, dec ; simpl.
|
||||||
destruct (isIn_decidable a {|Y & ϕ|}) as [t | t]
|
destruct (isIn_decidable a {|Y & ϕ|}) as [t | t]
|
||||||
; destruct (isIn_decidable a Y) as [n | n] ; rewrite comprehension_isIn in t
|
; destruct (isIn_decidable a Y) as [n | n] ; rewrite comprehension_isIn in t
|
||||||
; destruct (ϕ a) ; try reflexivity ; try contradiction.
|
; destruct (ϕ a) ; try reflexivity ; try contradiction
|
||||||
|
; try (contradiction (n t)) ; contradiction (t n).
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Lemma intersection_isIn_b (X Y: FSet A) (a : A) :
|
Lemma intersection_isIn_b (X Y: FSet A) (a : A) :
|
||||||
isIn_b a (intersection X Y) = andb (isIn_b a X) (isIn_b a Y).
|
a ∈_d (intersection X Y) = andb (a ∈_d X) (a ∈_d Y).
|
||||||
Proof.
|
Proof.
|
||||||
apply comprehension_isIn_b.
|
apply comprehension_isIn_b.
|
||||||
Defined.
|
Defined.
|
||||||
|
@ -114,7 +115,7 @@ Section SetLattice.
|
||||||
Instance fset_min : minimum (FSet A).
|
Instance fset_min : minimum (FSet A).
|
||||||
Proof.
|
Proof.
|
||||||
intros x y.
|
intros x y.
|
||||||
apply (intersection x y).
|
apply (x ∩ y).
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Instance fset_bot : bottom (FSet A).
|
Instance fset_bot : bottom (FSet A).
|
||||||
|
|
|
@ -48,31 +48,44 @@ Arguments neutralityR {_} {_} {_} {_} _.
|
||||||
Arguments absorb {_} {_} {_} {_} _ _.
|
Arguments absorb {_} {_} {_} {_} _ _.
|
||||||
|
|
||||||
Section structure.
|
Section structure.
|
||||||
Variable (T : Type -> Type).
|
Variable (T A : Type).
|
||||||
|
|
||||||
Class hasMembership : Type :=
|
Class hasMembership : Type :=
|
||||||
member : forall A : Type, A -> T A -> hProp.
|
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.
|
||||||
|
|
||||||
Class hasEmpty : Type :=
|
Class hasEmpty : Type :=
|
||||||
empty : forall A, T A.
|
empty : T.
|
||||||
|
|
||||||
Class hasSingleton : Type :=
|
Class hasSingleton : Type :=
|
||||||
singleton : forall A, A -> T A.
|
singleton : A -> T.
|
||||||
|
|
||||||
Class hasUnion : Type :=
|
Class hasUnion : Type :=
|
||||||
union : forall A, T A -> T A -> T A.
|
union : T -> T -> T.
|
||||||
|
|
||||||
Class hasIntersection : Type :=
|
Class hasIntersection : Type :=
|
||||||
intersection : forall A, T A -> T A -> T A.
|
intersection : T -> T -> T.
|
||||||
|
|
||||||
Class hasComprehension : Type :=
|
Class hasComprehension : Type :=
|
||||||
filter : forall A, (A -> Bool) -> T A -> T A.
|
filter : (A -> Bool) -> T -> T.
|
||||||
End structure.
|
End structure.
|
||||||
|
|
||||||
Arguments member {_} {_} {_} _ _.
|
Arguments member {_} {_} {_} _ _.
|
||||||
Arguments empty {_} {_} {_}.
|
Arguments subset {_} {_} _ _.
|
||||||
|
Arguments member_dec {_} {_} {_} _ _.
|
||||||
|
Arguments subset_dec {_} {_} _ _.
|
||||||
|
Arguments empty {_} {_}.
|
||||||
Arguments singleton {_} {_} {_} _.
|
Arguments singleton {_} {_} {_} _.
|
||||||
Arguments union {_} {_} {_} _ _.
|
Arguments union {_} {_} _ _.
|
||||||
|
Arguments intersection {_} {_} _ _.
|
||||||
Arguments filter {_} {_} {_} _ _.
|
Arguments filter {_} {_} {_} _ _.
|
||||||
|
|
||||||
Notation "∅" := empty.
|
Notation "∅" := empty.
|
||||||
|
@ -86,4 +99,7 @@ Notation "( ∩ )" := intersection (only parsing).
|
||||||
Notation "( X ∩ )" := (intersection X) (only parsing).
|
Notation "( X ∩ )" := (intersection X) (only parsing).
|
||||||
Notation "( ∩ Y )" := (fun X => X ∩ Y) (only parsing).
|
Notation "( ∩ Y )" := (fun X => X ∩ Y) (only parsing).
|
||||||
Notation "{| X & ϕ |}" := (filter ϕ X).
|
Notation "{| X & ϕ |}" := (filter ϕ X).
|
||||||
Infix "∈" := member (at level 9, right associativity).
|
Infix "∈" := member (at level 9, right associativity).
|
||||||
|
Infix "⊆" := subset (at level 10, right associativity).
|
||||||
|
Infix "∈_d" := member_dec (at level 9, right associativity).
|
||||||
|
Infix "⊆_d" := subset_dec (at level 10, right associativity).
|
|
@ -12,9 +12,9 @@ Module Export FSet.
|
||||||
| L : A -> FSet A
|
| L : A -> FSet A
|
||||||
| U : FSet A -> FSet A -> FSet A.
|
| U : FSet A -> FSet A -> FSet A.
|
||||||
|
|
||||||
Global Instance fset_empty : hasEmpty FSet := E.
|
Global Instance fset_empty : forall A, hasEmpty (FSet A) := E.
|
||||||
Global Instance fset_singleton : hasSingleton FSet := L.
|
Global Instance fset_singleton : forall A, hasSingleton (FSet A) A := L.
|
||||||
Global Instance fset_union : hasUnion FSet := U.
|
Global Instance fset_union : forall A, hasUnion (FSet A) := U.
|
||||||
|
|
||||||
Variable A : Type.
|
Variable A : Type.
|
||||||
|
|
||||||
|
|
|
@ -9,7 +9,7 @@ Module Export FSetC.
|
||||||
| Nil : FSetC A
|
| Nil : FSetC A
|
||||||
| Cns : A -> FSetC A -> FSetC A.
|
| Cns : A -> FSetC A -> FSetC A.
|
||||||
|
|
||||||
Global Instance fset_empty : hasEmpty FSetC := Nil.
|
Global Instance fset_empty : forall A,hasEmpty (FSetC A) := Nil.
|
||||||
|
|
||||||
Variable A : Type.
|
Variable A : Type.
|
||||||
Arguments Cns {_} _ _.
|
Arguments Cns {_} _ _.
|
||||||
|
|
|
@ -9,9 +9,9 @@ Module Export FSet.
|
||||||
| L : A -> FSet A
|
| L : A -> FSet A
|
||||||
| U : FSet A -> FSet A -> FSet A.
|
| U : FSet A -> FSet A -> FSet A.
|
||||||
|
|
||||||
Global Instance fset_empty : hasEmpty FSet := E.
|
Global Instance fset_empty : forall A, hasEmpty (FSet A) := E.
|
||||||
Global Instance fset_singleton : hasSingleton FSet := L.
|
Global Instance fset_singleton : forall A, hasSingleton (FSet A) A := L.
|
||||||
Global Instance fset_union : hasUnion FSet := U.
|
Global Instance fset_union : forall A, hasUnion (FSet A) := U.
|
||||||
|
|
||||||
Variable A : Type.
|
Variable A : Type.
|
||||||
|
|
||||||
|
|
Loading…
Reference in New Issue