mirror of
https://github.com/nmvdw/HITs-Examples
synced 2025-11-03 07:03:51 +01:00
A negligible change in the structure
This commit is contained in:
250
FiniteSets/interfaces/lattice_examples.v
Normal file
250
FiniteSets/interfaces/lattice_examples.v
Normal file
@@ -0,0 +1,250 @@
|
||||
(** Some examples of lattices. *)
|
||||
Require Import HoTT lattice_interface.
|
||||
|
||||
(** [Bool] is a lattice. *)
|
||||
Section BoolLattice.
|
||||
Ltac solve_bool :=
|
||||
let x := fresh in
|
||||
repeat (intro x ; destruct x)
|
||||
; compute
|
||||
; auto
|
||||
; try contradiction.
|
||||
|
||||
Instance maximum_bool : maximum Bool := orb.
|
||||
Instance minimum_bool : minimum Bool := andb.
|
||||
Instance bottom_bool : bottom Bool := false.
|
||||
|
||||
Global Instance lattice_bool : Lattice Bool.
|
||||
Proof.
|
||||
split ; solve_bool.
|
||||
Defined.
|
||||
|
||||
Definition and_true : forall b, andb b true = b.
|
||||
Proof.
|
||||
solve_bool.
|
||||
Defined.
|
||||
|
||||
Definition and_false : forall b, andb b false = false.
|
||||
Proof.
|
||||
solve_bool.
|
||||
Defined.
|
||||
|
||||
Definition dist₁ : forall b₁ b₂ b₃,
|
||||
andb b₁ (orb b₂ b₃) = orb (andb b₁ b₂) (andb b₁ b₃).
|
||||
Proof.
|
||||
solve_bool.
|
||||
Defined.
|
||||
|
||||
Definition dist₂ : forall b₁ b₂ b₃,
|
||||
orb b₁ (andb b₂ b₃) = andb (orb b₁ b₂) (orb b₁ b₃).
|
||||
Proof.
|
||||
solve_bool.
|
||||
Defined.
|
||||
|
||||
Definition max_min : forall b₁ b₂,
|
||||
orb (andb b₁ b₂) b₁ = b₁.
|
||||
Proof.
|
||||
solve_bool.
|
||||
Defined.
|
||||
End BoolLattice.
|
||||
|
||||
Create HintDb bool_lattice_hints.
|
||||
Hint Resolve associativity : bool_lattice_hints.
|
||||
Hint Resolve (associativity _ _ _)^ : bool_lattice_hints.
|
||||
Hint Resolve commutative : bool_lattice_hints.
|
||||
Hint Resolve absorb : bool_lattice_hints.
|
||||
Hint Resolve idempotency : bool_lattice_hints.
|
||||
Hint Resolve neutralityL : bool_lattice_hints.
|
||||
Hint Resolve neutralityR : bool_lattice_hints.
|
||||
|
||||
Hint Resolve
|
||||
associativity
|
||||
and_true and_false
|
||||
dist₁ dist₂ max_min
|
||||
: bool_lattice_hints.
|
||||
|
||||
(** If [B] is a lattice, then [A -> B] is a lattice. *)
|
||||
Section fun_lattice.
|
||||
Context {A B : Type}.
|
||||
Context `{Lattice B}.
|
||||
Context `{Funext}.
|
||||
|
||||
Global Instance max_fun : maximum (A -> B) :=
|
||||
fun (f g : A -> B) (a : A) => max_L0 (f a) (g a).
|
||||
|
||||
Global Instance min_fun : minimum (A -> B) :=
|
||||
fun (f g : A -> B) (a : A) => min_L0 (f a) (g a).
|
||||
|
||||
Global Instance bot_fun : bottom (A -> B)
|
||||
:= fun _ => empty_L.
|
||||
|
||||
Ltac solve_fun :=
|
||||
compute ; intros ; apply path_forall ; intro ;
|
||||
eauto with lattice_hints typeclass_instances.
|
||||
|
||||
Global Instance lattice_fun : Lattice (A -> B).
|
||||
Proof.
|
||||
split ; solve_fun.
|
||||
Defined.
|
||||
End fun_lattice.
|
||||
|
||||
(** If [A] is a lattice and [P] is closed under the lattice operations, then [Σ(x:A), P x] is a lattice. *)
|
||||
Section sub_lattice.
|
||||
Context {A : Type} {P : A -> hProp}.
|
||||
Context `{Lattice A}.
|
||||
Context {Hmax : forall x y, P x -> P y -> P (max_L0 x y)}.
|
||||
Context {Hmin : forall x y, P x -> P y -> P (min_L0 x y)}.
|
||||
Context {Hbot : P empty_L}.
|
||||
|
||||
Definition AP : Type := sig P.
|
||||
|
||||
Instance botAP : bottom AP := (empty_L ; Hbot).
|
||||
|
||||
Instance maxAP : maximum AP :=
|
||||
fun x y =>
|
||||
match x, y with
|
||||
| (a ; pa), (b ; pb) => (max_L0 a b ; Hmax a b pa pb)
|
||||
end.
|
||||
|
||||
Instance minAP : minimum AP :=
|
||||
fun x y =>
|
||||
match x, y with
|
||||
| (a ; pa), (b ; pb) => (min_L0 a b ; Hmin a b pa pb)
|
||||
end.
|
||||
|
||||
Instance hprop_sub : forall c, IsHProp (P c).
|
||||
Proof.
|
||||
apply _.
|
||||
Defined.
|
||||
|
||||
Ltac solve_sub :=
|
||||
let x := fresh in
|
||||
repeat (intro x ; destruct x)
|
||||
; simple refine (path_sigma _ _ _ _ _)
|
||||
; simpl
|
||||
; try (apply hprop_sub)
|
||||
; eauto 3 with lattice_hints typeclass_instances.
|
||||
|
||||
Global Instance lattice_sub : Lattice AP.
|
||||
Proof.
|
||||
split ; solve_sub.
|
||||
Defined.
|
||||
End sub_lattice.
|
||||
|
||||
Instance lor : maximum hProp := fun X Y => BuildhProp (Trunc (-1) (sum X Y)).
|
||||
|
||||
Delimit Scope logic_scope with L.
|
||||
Notation "A ∨ B" := (lor A B) (at level 20, right associativity) : logic_scope.
|
||||
Arguments lor _%L _%L.
|
||||
Open Scope logic_scope.
|
||||
|
||||
Instance land : minimum hProp := fun X Y => BuildhProp (prod X Y).
|
||||
Instance lfalse : bottom hProp := False_hp.
|
||||
|
||||
Notation "A ∧ B" := (land A B) (at level 20, right associativity) : logic_scope.
|
||||
Arguments land _%L _%L.
|
||||
Open Scope logic_scope.
|
||||
|
||||
(** [hProp] is a lattice. *)
|
||||
Section hPropLattice.
|
||||
Context `{Univalence}.
|
||||
|
||||
Local Ltac lor_intros :=
|
||||
let x := fresh in intro x
|
||||
; repeat (strip_truncations ; destruct x as [x | x]).
|
||||
|
||||
Instance lor_commutative : Commutative lor.
|
||||
Proof.
|
||||
intros X Y.
|
||||
apply path_iff_hprop ; lor_intros
|
||||
; apply tr ; auto.
|
||||
Defined.
|
||||
|
||||
Instance land_commutative : Commutative land.
|
||||
Proof.
|
||||
intros X Y.
|
||||
apply path_hprop.
|
||||
apply equiv_prod_symm.
|
||||
Defined.
|
||||
|
||||
Instance lor_associative : Associative lor.
|
||||
Proof.
|
||||
intros X Y Z.
|
||||
apply path_iff_hprop ; lor_intros
|
||||
; apply tr ; auto
|
||||
; try (left ; apply tr)
|
||||
; try (right ; apply tr) ; auto.
|
||||
Defined.
|
||||
|
||||
Instance land_associative : Associative land.
|
||||
Proof.
|
||||
intros X Y Z.
|
||||
symmetry.
|
||||
apply path_hprop.
|
||||
apply equiv_prod_assoc.
|
||||
Defined.
|
||||
|
||||
Instance lor_idempotent : Idempotent lor.
|
||||
Proof.
|
||||
intros X.
|
||||
apply path_iff_hprop ; lor_intros
|
||||
; try(refine (tr(inl _))) ; auto.
|
||||
Defined.
|
||||
|
||||
Instance land_idempotent : Idempotent land.
|
||||
Proof.
|
||||
intros X.
|
||||
apply path_iff_hprop ; cbn.
|
||||
- intros [a b] ; apply a.
|
||||
- intros a ; apply (pair a a).
|
||||
Defined.
|
||||
|
||||
Instance lor_neutrall : NeutralL lor lfalse.
|
||||
Proof.
|
||||
intros X.
|
||||
apply path_iff_hprop ; lor_intros ; try contradiction
|
||||
; try (refine (tr(inr _))) ; auto.
|
||||
Defined.
|
||||
|
||||
Instance lor_neutralr : NeutralR lor lfalse.
|
||||
Proof.
|
||||
intros X.
|
||||
apply path_iff_hprop ; lor_intros ; try contradiction
|
||||
; try (refine (tr(inl _))) ; auto.
|
||||
Defined.
|
||||
|
||||
Instance absorption_orb_andb : Absorption lor land.
|
||||
Proof.
|
||||
intros Z1 Z2.
|
||||
apply path_iff_hprop ; cbn.
|
||||
- intros X ; strip_truncations.
|
||||
destruct X as [Xx | [Xy1 Xy2]] ; assumption.
|
||||
- intros X.
|
||||
apply (tr (inl X)).
|
||||
Defined.
|
||||
|
||||
Instance absorption_andb_orb : Absorption land lor.
|
||||
Proof.
|
||||
intros Z1 Z2.
|
||||
apply path_iff_hprop ; cbn.
|
||||
- intros [X Y] ; strip_truncations.
|
||||
assumption.
|
||||
- intros X.
|
||||
split.
|
||||
* assumption.
|
||||
* apply (tr (inl X)).
|
||||
Defined.
|
||||
|
||||
Global Instance lattice_hprop : Lattice hProp :=
|
||||
{ commutative_min := _ ;
|
||||
commutative_max := _ ;
|
||||
associative_min := _ ;
|
||||
associative_max := _ ;
|
||||
idempotent_min := _ ;
|
||||
idempotent_max := _ ;
|
||||
neutralL_max := _ ;
|
||||
neutralR_max := _ ;
|
||||
absorption_min_max := _ ;
|
||||
absorption_max_min := _
|
||||
}.
|
||||
End hPropLattice.
|
||||
115
FiniteSets/interfaces/lattice_interface.v
Normal file
115
FiniteSets/interfaces/lattice_interface.v
Normal file
@@ -0,0 +1,115 @@
|
||||
(** Interface for lattices and join semilattices. *)
|
||||
Require Import HoTT.
|
||||
|
||||
Section binary_operation.
|
||||
Definition operation (A : Type) := A -> A -> A.
|
||||
|
||||
Variable (A : Type)
|
||||
(f : operation A).
|
||||
|
||||
Class Commutative :=
|
||||
commutative : forall x y, f x y = f y x.
|
||||
|
||||
Class Associative :=
|
||||
associativity : forall x y z, f (f x y) z = f x (f y z).
|
||||
|
||||
Class Idempotent :=
|
||||
idempotency : forall x, f x x = x.
|
||||
|
||||
Variable g : operation A.
|
||||
|
||||
Class Absorption :=
|
||||
absorb : forall x y, f x (g x y) = x.
|
||||
|
||||
Variable (n : A).
|
||||
|
||||
Class NeutralL :=
|
||||
neutralityL : forall x, f n x = x.
|
||||
|
||||
Class NeutralR :=
|
||||
neutralityR : forall x, f x n = x.
|
||||
End binary_operation.
|
||||
|
||||
Arguments Commutative {_} _.
|
||||
Arguments Associative {_} _.
|
||||
Arguments Idempotent {_} _.
|
||||
Arguments NeutralL {_} _ _.
|
||||
Arguments NeutralR {_} _ _.
|
||||
Arguments Absorption {_} _ _.
|
||||
Arguments commutative {_} {_} {_} _ _.
|
||||
Arguments associativity {_} {_} {_} _ _ _.
|
||||
Arguments idempotency {_} {_} {_} _.
|
||||
Arguments neutralityL {_} {_} {_} {_} _.
|
||||
Arguments neutralityR {_} {_} {_} {_} _.
|
||||
Arguments absorb {_} {_} {_} {_} _ _.
|
||||
|
||||
Section lattice_operations.
|
||||
Variable (A : Type).
|
||||
|
||||
Class maximum :=
|
||||
max_L : operation A.
|
||||
|
||||
Class minimum :=
|
||||
min_L : operation A.
|
||||
|
||||
Class bottom :=
|
||||
empty : A.
|
||||
End lattice_operations.
|
||||
|
||||
Arguments max_L {_} {_} _.
|
||||
Arguments min_L {_} {_} _.
|
||||
Arguments empty {_}.
|
||||
|
||||
Section JoinSemiLattice.
|
||||
Variable A : Type.
|
||||
Context {max_L : maximum A} {empty_L : bottom A}.
|
||||
|
||||
Class JoinSemiLattice :=
|
||||
{
|
||||
commutative_max_js :> Commutative max_L ;
|
||||
associative_max_js :> Associative max_L ;
|
||||
idempotent_max_js :> Idempotent max_L ;
|
||||
neutralL_max_js :> NeutralL max_L empty_L ;
|
||||
neutralR_max_js :> NeutralR max_L empty_L ;
|
||||
}.
|
||||
End JoinSemiLattice.
|
||||
|
||||
Arguments JoinSemiLattice _ {_} {_}.
|
||||
|
||||
Create HintDb joinsemilattice_hints.
|
||||
Hint Resolve associativity : joinsemilattice_hints.
|
||||
Hint Resolve (associativity _ _ _)^ : joinsemilattice_hints.
|
||||
Hint Resolve commutative : joinsemilattice_hints.
|
||||
Hint Resolve idempotency : joinsemilattice_hints.
|
||||
Hint Resolve neutralityL : joinsemilattice_hints.
|
||||
Hint Resolve neutralityR : joinsemilattice_hints.
|
||||
|
||||
Section Lattice.
|
||||
Variable A : Type.
|
||||
Context {max_L : maximum A} {min_L : minimum A} {empty_L : bottom A}.
|
||||
|
||||
Class Lattice :=
|
||||
{
|
||||
commutative_min :> Commutative min_L ;
|
||||
commutative_max :> Commutative max_L ;
|
||||
associative_min :> Associative min_L ;
|
||||
associative_max :> Associative max_L ;
|
||||
idempotent_min :> Idempotent min_L ;
|
||||
idempotent_max :> Idempotent max_L ;
|
||||
neutralL_max :> NeutralL max_L empty_L ;
|
||||
neutralR_max :> NeutralR max_L empty_L ;
|
||||
absorption_min_max :> Absorption min_L max_L ;
|
||||
absorption_max_min :> Absorption max_L min_L
|
||||
}.
|
||||
End Lattice.
|
||||
|
||||
Arguments Lattice _ {_} {_} {_}.
|
||||
|
||||
Create HintDb lattice_hints.
|
||||
Hint Resolve associativity : lattice_hints.
|
||||
Hint Resolve (associativity _ _ _)^ : lattice_hints.
|
||||
Hint Resolve commutative : lattice_hints.
|
||||
Hint Resolve absorb : lattice_hints.
|
||||
Hint Resolve idempotency : lattice_hints.
|
||||
Hint Resolve neutralityL : lattice_hints.
|
||||
Hint Resolve neutralityR : lattice_hints.
|
||||
45
FiniteSets/interfaces/monad_interface.v
Normal file
45
FiniteSets/interfaces/monad_interface.v
Normal file
@@ -0,0 +1,45 @@
|
||||
(** The structure of a monad. *)
|
||||
Require Import HoTT.
|
||||
|
||||
Section functor.
|
||||
Variable (F : Type -> Type).
|
||||
Context `{Functorish F}.
|
||||
|
||||
Class Functor :=
|
||||
{
|
||||
fmap_compose {A B C : Type} (f : A -> B) (g : B -> C) :
|
||||
fmap F (g o f) = fmap F g o fmap F f
|
||||
}.
|
||||
End functor.
|
||||
|
||||
Section monad_operations.
|
||||
Variable (F : Type -> Type).
|
||||
|
||||
Class hasPure :=
|
||||
{
|
||||
pure : forall (A : Type), A -> F A
|
||||
}.
|
||||
|
||||
Class hasBind :=
|
||||
{
|
||||
bind : forall (A : Type), F(F A) -> F A
|
||||
}.
|
||||
End monad_operations.
|
||||
|
||||
Arguments pure {_} {_} _ _.
|
||||
Arguments bind {_} {_} _ _.
|
||||
|
||||
Section monad.
|
||||
Variable (F : Type -> Type).
|
||||
Context `{Functor F} `{hasPure F} `{hasBind F}.
|
||||
|
||||
Class Monad :=
|
||||
{
|
||||
bind_assoc : forall {A : Type},
|
||||
bind A o bind (F A) = bind A o fmap F (bind A) ;
|
||||
bind_neutral_l : forall {A : Type},
|
||||
bind A o pure (F A) = idmap ;
|
||||
bind_neutral_r : forall {A : Type},
|
||||
bind A o fmap F (pure A) = idmap
|
||||
}.
|
||||
End monad.
|
||||
331
FiniteSets/interfaces/set_interface.v
Normal file
331
FiniteSets/interfaces/set_interface.v
Normal file
@@ -0,0 +1,331 @@
|
||||
Require Import HoTT.
|
||||
Require Import FSets lattice_interface.
|
||||
|
||||
Section interface.
|
||||
Context `{Univalence}.
|
||||
Variable (T : Type -> Type)
|
||||
(f : forall A, T A -> FSet A).
|
||||
Context `{forall A, hasMembership (T A) A
|
||||
, forall A, hasEmpty (T A)
|
||||
, forall A, hasSingleton (T A) A
|
||||
, forall A, hasUnion (T A)
|
||||
, forall A, hasComprehension (T A) A}.
|
||||
|
||||
Class sets :=
|
||||
{
|
||||
f_empty : forall A, f A ∅ = ∅ ;
|
||||
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) = {| f A X & φ |};
|
||||
f_member : forall A a X, member a X = a ∈ (f A X)
|
||||
}.
|
||||
|
||||
Global Instance f_surjective A `{sets} : IsSurjection (f A).
|
||||
Proof.
|
||||
unfold IsSurjection.
|
||||
hinduction ; try (intros ; apply path_ishprop).
|
||||
- simple refine (BuildContr _ _ _).
|
||||
* refine (tr(∅;_)).
|
||||
apply f_empty.
|
||||
* intros ; apply path_ishprop.
|
||||
- intro a.
|
||||
simple refine (BuildContr _ _ _).
|
||||
* refine (tr({|a|};_)).
|
||||
apply f_singleton.
|
||||
* intros ; apply path_ishprop.
|
||||
- intros Y1 Y2 [X1' ?] [X2' ?].
|
||||
simple refine (BuildContr _ _ _).
|
||||
* simple refine (Trunc_rec _ X1') ; intros [X1 fX1].
|
||||
simple refine (Trunc_rec _ X2') ; intros [X2 fX2].
|
||||
refine (tr(X1 ∪ X2;_)).
|
||||
rewrite f_union, fX1, fX2.
|
||||
reflexivity.
|
||||
* intros ; apply path_ishprop.
|
||||
Defined.
|
||||
|
||||
End interface.
|
||||
|
||||
Section quotient_surjection.
|
||||
Variable (A B : Type)
|
||||
(f : A -> B)
|
||||
(H : IsSurjection f).
|
||||
Context `{IsHSet B} `{Univalence}.
|
||||
|
||||
Definition f_eq : relation A := fun a1 a2 => f a1 = f a2.
|
||||
Definition quotientB : Type := quotient f_eq.
|
||||
|
||||
Global Instance quotientB_recursion : HitRecursion quotientB :=
|
||||
{
|
||||
indTy := _;
|
||||
recTy :=
|
||||
forall (P : Type) (HP: IsHSet P) (u : A -> P),
|
||||
(forall x y : A, f_eq x y -> u x = u y) -> quotientB -> P;
|
||||
H_inductor := quotient_ind f_eq ;
|
||||
H_recursor := @quotient_rec _ f_eq _
|
||||
}.
|
||||
|
||||
Global Instance R_refl : Reflexive f_eq.
|
||||
Proof.
|
||||
intro.
|
||||
reflexivity.
|
||||
Defined.
|
||||
|
||||
Global Instance R_sym : Symmetric f_eq.
|
||||
Proof.
|
||||
intros a b Hab.
|
||||
apply (Hab^).
|
||||
Defined.
|
||||
|
||||
Global Instance R_trans : Transitive f_eq.
|
||||
Proof.
|
||||
intros a b c Hab Hbc.
|
||||
apply (Hab @ Hbc).
|
||||
Defined.
|
||||
|
||||
Definition quotientB_to_B : quotientB -> B.
|
||||
Proof.
|
||||
hrecursion.
|
||||
- apply f.
|
||||
- done.
|
||||
Defined.
|
||||
|
||||
Instance quotientB_emb : IsEmbedding (quotientB_to_B).
|
||||
Proof.
|
||||
apply isembedding_isinj_hset.
|
||||
unfold isinj.
|
||||
hrecursion ; [ | intros; apply path_ishprop ].
|
||||
intro.
|
||||
hrecursion ; [ | intros; apply path_ishprop ].
|
||||
intros.
|
||||
by apply related_classes_eq.
|
||||
Defined.
|
||||
|
||||
Instance quotientB_surj : IsSurjection (quotientB_to_B).
|
||||
Proof.
|
||||
apply BuildIsSurjection.
|
||||
intros Y.
|
||||
destruct (H Y).
|
||||
simple refine (Trunc_rec _ center) ; intros [a fa].
|
||||
apply (tr(class_of _ a;fa)).
|
||||
Defined.
|
||||
|
||||
Definition quotient_iso : quotientB <~> B.
|
||||
Proof.
|
||||
refine (BuildEquiv _ _ quotientB_to_B _).
|
||||
apply isequiv_surj_emb ; apply _.
|
||||
Defined.
|
||||
|
||||
Definition reflect_eq : forall (X Y : A),
|
||||
f X = f Y -> f_eq X Y.
|
||||
Proof.
|
||||
done.
|
||||
Defined.
|
||||
|
||||
Lemma same_class : forall (X Y : A),
|
||||
class_of f_eq X = class_of f_eq Y -> f_eq X Y.
|
||||
Proof.
|
||||
intros.
|
||||
simple refine (classes_eq_related _ _ _ _) ; assumption.
|
||||
Defined.
|
||||
|
||||
End quotient_surjection.
|
||||
|
||||
Arguments quotient_iso {_} {_} _ {_} {_} {_}.
|
||||
|
||||
Ltac reduce T :=
|
||||
intros ;
|
||||
repeat (rewrite (f_empty T _)
|
||||
|| rewrite (f_singleton T _)
|
||||
|| rewrite (f_union T _)
|
||||
|| rewrite (f_filter T _)
|
||||
|| rewrite (f_member T _)).
|
||||
|
||||
Section quotient_properties.
|
||||
Variable (T : Type -> Type).
|
||||
Variable (f : forall {A : Type}, T A -> FSet A).
|
||||
Context `{sets T f}.
|
||||
|
||||
Definition set_eq A := f_eq (T A) (FSet A) (f A).
|
||||
Definition View A : Type := quotientB (T A) (FSet A) (f A).
|
||||
|
||||
Instance f_is_surjective A : IsSurjection (f A).
|
||||
Proof.
|
||||
apply (f_surjective T f A).
|
||||
Defined.
|
||||
|
||||
Global Instance view_union (A : Type) : hasUnion (View A).
|
||||
Proof.
|
||||
intros X Y.
|
||||
apply (quotient_iso _)^-1.
|
||||
simple refine (union _ _).
|
||||
- simple refine (quotient_iso (f A) X).
|
||||
- simple refine (quotient_iso (f A) Y).
|
||||
Defined.
|
||||
|
||||
Definition well_defined_union (A : Type) (X Y : T A) :
|
||||
(class_of _ X) ∪ (class_of _ Y) = class_of _ (X ∪ Y).
|
||||
Proof.
|
||||
rewrite <- (eissect (quotient_iso _)).
|
||||
simpl.
|
||||
rewrite (f_union T _).
|
||||
reflexivity.
|
||||
Defined.
|
||||
|
||||
Global Instance view_comprehension (A : Type) : hasComprehension (View A) A.
|
||||
Proof.
|
||||
intros ϕ X.
|
||||
apply (quotient_iso _)^-1.
|
||||
simple refine ({|_ & ϕ|}).
|
||||
apply (quotient_iso (f A) X).
|
||||
Defined.
|
||||
|
||||
Definition well_defined_filter (A : Type) (ϕ : A -> Bool) (X : T A) :
|
||||
{|class_of _ X & ϕ|} = class_of _ {|X & ϕ|}.
|
||||
Proof.
|
||||
rewrite <- (eissect (quotient_iso _)).
|
||||
simpl.
|
||||
rewrite (f_filter T _).
|
||||
reflexivity.
|
||||
Defined.
|
||||
|
||||
Global Instance View_empty A : hasEmpty (View A).
|
||||
Proof.
|
||||
apply ((quotient_iso _)^-1 ∅).
|
||||
Defined.
|
||||
|
||||
Definition well_defined_empty A : ∅ = class_of (set_eq A) ∅.
|
||||
Proof.
|
||||
rewrite <- (eissect (quotient_iso _)).
|
||||
simpl.
|
||||
rewrite (f_empty T _).
|
||||
reflexivity.
|
||||
Defined.
|
||||
|
||||
Global Instance View_singleton A: hasSingleton (View A) A.
|
||||
Proof.
|
||||
intro a ; apply ((quotient_iso _)^-1 {|a|}).
|
||||
Defined.
|
||||
|
||||
Definition well_defined_sungleton A (a : A) : {|a|} = class_of _ {|a|}.
|
||||
Proof.
|
||||
rewrite <- (eissect (quotient_iso _)).
|
||||
simpl.
|
||||
rewrite (f_singleton T _).
|
||||
reflexivity.
|
||||
Defined.
|
||||
|
||||
Global Instance View_member A : hasMembership (View A) A.
|
||||
Proof.
|
||||
intros a ; unfold View.
|
||||
hrecursion.
|
||||
- apply (member a).
|
||||
- intros X Y HXY.
|
||||
reduce T.
|
||||
apply (ap _ HXY).
|
||||
Defined.
|
||||
|
||||
Instance View_max A : maximum (View A).
|
||||
Proof.
|
||||
apply view_union.
|
||||
Defined.
|
||||
|
||||
Hint Unfold Commutative Associative Idempotent NeutralL NeutralR View_max view_union.
|
||||
|
||||
Instance bottom_view A : bottom (View A).
|
||||
Proof.
|
||||
apply View_empty.
|
||||
Defined.
|
||||
|
||||
Ltac sq1 := autounfold ; intros ; try f_ap
|
||||
; rewrite ?(eisretr (quotient_iso _))
|
||||
; eauto with lattice_hints typeclass_instances.
|
||||
|
||||
Ltac sq2 := autounfold ; intros
|
||||
; rewrite <- (eissect (quotient_iso _)), ?(eisretr (quotient_iso _))
|
||||
; f_ap ; simpl
|
||||
; reduce T ; eauto with lattice_hints typeclass_instances.
|
||||
|
||||
Global Instance view_lattice A : JoinSemiLattice (View A).
|
||||
Proof.
|
||||
split ; try sq1 ; try sq2.
|
||||
Defined.
|
||||
|
||||
End quotient_properties.
|
||||
|
||||
Arguments set_eq {_} _ {_} _ _.
|
||||
|
||||
Section properties.
|
||||
Context `{Univalence}.
|
||||
Variable (T : Type -> Type) (f : forall A, T A -> FSet A).
|
||||
Context `{sets T f}.
|
||||
|
||||
Definition set_subset : forall A, T A -> T A -> hProp :=
|
||||
fun A X Y => (f A X) ⊆ (f A Y).
|
||||
|
||||
Definition empty_isIn : forall (A : Type) (a : A),
|
||||
a ∈ ∅ = False_hp.
|
||||
Proof.
|
||||
by (reduce T).
|
||||
Defined.
|
||||
|
||||
Definition singleton_isIn : forall (A : Type) (a b : A),
|
||||
a ∈ {|b|} = merely (a = b).
|
||||
Proof.
|
||||
by (reduce T).
|
||||
Defined.
|
||||
|
||||
Definition union_isIn : forall (A : Type) (a : A) (X Y : T A),
|
||||
a ∈ (X ∪ Y) = lor (a ∈ X) (a ∈ Y).
|
||||
Proof.
|
||||
by (reduce T).
|
||||
Defined.
|
||||
|
||||
Definition filter_isIn : forall (A : Type) (a : A) (ϕ : A -> Bool) (X : T A),
|
||||
member a (filter ϕ X) = if ϕ a then member a X else False_hp.
|
||||
Proof.
|
||||
reduce T.
|
||||
apply properties.comprehension_isIn.
|
||||
Defined.
|
||||
|
||||
Definition reflect_f_eq : forall (A : Type) (X Y : T A),
|
||||
class_of (set_eq f) X = class_of (set_eq f) Y -> set_eq f X Y.
|
||||
Proof.
|
||||
intros.
|
||||
refine (same_class _ _ _ _ _ _) ; assumption.
|
||||
Defined.
|
||||
|
||||
Ltac via_quotient := intros ; apply reflect_f_eq ; simpl
|
||||
; rewrite <- ?(well_defined_union T _), <- ?(well_defined_empty T _)
|
||||
; eauto with lattice_hints typeclass_instances.
|
||||
|
||||
Lemma union_comm : forall A (X Y : T A),
|
||||
set_eq f (X ∪ Y) (Y ∪ X).
|
||||
Proof.
|
||||
via_quotient.
|
||||
Defined.
|
||||
|
||||
Lemma union_assoc : forall A (X Y Z : T A),
|
||||
set_eq f ((X ∪ Y) ∪ Z) (X ∪ (Y ∪ Z)).
|
||||
Proof.
|
||||
via_quotient.
|
||||
Defined.
|
||||
|
||||
Lemma union_idem : forall A (X : T A),
|
||||
set_eq f (X ∪ X) X.
|
||||
Proof.
|
||||
via_quotient.
|
||||
Defined.
|
||||
|
||||
Lemma union_neutralL : forall A (X : T A),
|
||||
set_eq f (∅ ∪ X) X.
|
||||
Proof.
|
||||
via_quotient.
|
||||
Defined.
|
||||
|
||||
Lemma union_neutralR : forall A (X : T A),
|
||||
set_eq f (X ∪ ∅) X.
|
||||
Proof.
|
||||
via_quotient.
|
||||
Defined.
|
||||
|
||||
End properties.
|
||||
59
FiniteSets/interfaces/set_names.v
Normal file
59
FiniteSets/interfaces/set_names.v
Normal file
@@ -0,0 +1,59 @@
|
||||
(** Classes for set operations, so they can be overloaded. *)
|
||||
Require Import HoTT.
|
||||
|
||||
Section structure.
|
||||
Variable (T A : Type).
|
||||
|
||||
Class hasMembership : Type :=
|
||||
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 :=
|
||||
empty : T.
|
||||
|
||||
Class hasSingleton : Type :=
|
||||
singleton : A -> T.
|
||||
|
||||
Class hasUnion : Type :=
|
||||
union : T -> T -> T.
|
||||
|
||||
Class hasIntersection : Type :=
|
||||
intersection : T -> T -> T.
|
||||
|
||||
Class hasComprehension : Type :=
|
||||
filter : (A -> Bool) -> T -> T.
|
||||
End structure.
|
||||
|
||||
Arguments member {_} {_} {_} _ _.
|
||||
Arguments subset {_} {_} _ _.
|
||||
Arguments member_dec {_} {_} {_} _ _.
|
||||
Arguments subset_dec {_} {_} _ _.
|
||||
Arguments empty {_} {_}.
|
||||
Arguments singleton {_} {_} {_} _.
|
||||
Arguments union {_} {_} _ _.
|
||||
Arguments intersection {_} {_} _ _.
|
||||
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).
|
||||
Notation "{| X & ϕ |}" := (filter ϕ X).
|
||||
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).
|
||||
Reference in New Issue
Block a user