mirror of
https://github.com/nmvdw/HITs-Examples
synced 2025-11-04 07:33:51 +01:00
Port the codebase to HottClasses
Initial work + use the latest version of HoTT
This commit is contained in:
@@ -1,5 +1,5 @@
|
||||
(** Some examples of lattices. *)
|
||||
Require Import HoTT lattice_interface.
|
||||
Require Export HoTT lattice_interface.
|
||||
|
||||
(** [Bool] is a lattice. *)
|
||||
Section BoolLattice.
|
||||
@@ -10,14 +10,18 @@ Section BoolLattice.
|
||||
; auto
|
||||
; try contradiction.
|
||||
|
||||
Instance maximum_bool : maximum Bool := orb.
|
||||
Instance minimum_bool : minimum Bool := andb.
|
||||
Instance bottom_bool : bottom Bool := false.
|
||||
Instance maximum_bool : Join Bool := orb.
|
||||
Instance minimum_bool : Meet Bool := andb.
|
||||
Instance bottom_bool : Bottom Bool := false.
|
||||
|
||||
Global Instance boundedjoinsemilattice_bool : BoundedJoinSemiLattice Bool.
|
||||
Proof. repeat (split ; (apply _ || solve_bool)). Defined.
|
||||
Global Instance meetsemilattice_bool : MeetSemiLattice Bool.
|
||||
Proof. repeat (split ; (apply _ || solve_bool)). Defined.
|
||||
Global Instance boundedmeetsemilattice_bool : @BoundedSemiLattice Bool (⊓) true.
|
||||
Proof. repeat (split ; (apply _ || solve_bool)). Defined.
|
||||
Global Instance lattice_bool : Lattice Bool.
|
||||
Proof.
|
||||
split ; solve_bool.
|
||||
Defined.
|
||||
Proof. split ; (apply _ || solve_bool). Defined.
|
||||
|
||||
Definition and_true : forall b, andb b true = b.
|
||||
Proof.
|
||||
@@ -50,12 +54,12 @@ End BoolLattice.
|
||||
|
||||
Create HintDb bool_lattice_hints.
|
||||
Hint Resolve associativity : bool_lattice_hints.
|
||||
Hint Resolve (associativity _ _ _)^ : bool_lattice_hints.
|
||||
(* Hint Resolve (associativity _ _ _)^ : bool_lattice_hints. *)
|
||||
Hint Resolve commutativity : bool_lattice_hints.
|
||||
Hint Resolve absorb : bool_lattice_hints.
|
||||
Hint Resolve absorption : bool_lattice_hints.
|
||||
Hint Resolve idempotency : bool_lattice_hints.
|
||||
Hint Resolve neutralityL : bool_lattice_hints.
|
||||
Hint Resolve neutralityR : bool_lattice_hints.
|
||||
Hint Resolve left_identity : bool_lattice_hints.
|
||||
Hint Resolve right_identity : bool_lattice_hints.
|
||||
|
||||
Hint Resolve
|
||||
associativity
|
||||
@@ -66,25 +70,51 @@ Hint Resolve
|
||||
(** If [B] is a lattice, then [A -> B] is a lattice. *)
|
||||
Section fun_lattice.
|
||||
Context {A B : Type}.
|
||||
Context `{Lattice B}.
|
||||
Context `{BJoin : Join B}.
|
||||
Context `{BMeet : Meet B}.
|
||||
Context `{@Lattice B BJoin BMeet}.
|
||||
Context `{Funext}.
|
||||
Context `{BBottom : Bottom B}.
|
||||
|
||||
Global Instance max_fun : maximum (A -> B) :=
|
||||
fun (f g : A -> B) (a : A) => max_L0 (f a) (g a).
|
||||
Global Instance bot_fun : Bottom (A -> B)
|
||||
:= fun _ => ⊥.
|
||||
|
||||
Global Instance min_fun : minimum (A -> B) :=
|
||||
fun (f g : A -> B) (a : A) => min_L0 (f a) (g a).
|
||||
Global Instance join_fun : Join (A -> B) :=
|
||||
fun (f g : A -> B) (a : A) => (f a) ⊔ (g a).
|
||||
|
||||
Global Instance bot_fun : bottom (A -> B)
|
||||
:= fun _ => empty_L.
|
||||
Global Instance meet_fun : Meet (A -> B) :=
|
||||
fun (f g : A -> B) (a : A) => (f a) ⊓ (g a).
|
||||
|
||||
Ltac solve_fun :=
|
||||
compute ; intros ; apply path_forall ; intro ;
|
||||
eauto with lattice_hints typeclass_instances.
|
||||
|
||||
Create HintDb test1.
|
||||
Lemma associativity_lat `{Lattice A} (x y z : A) :
|
||||
x ⊓ (y ⊓ z) = x ⊓ y ⊓ z.
|
||||
Proof. apply associativity. Defined.
|
||||
Hint Resolve associativity : test1.
|
||||
Hint Resolve associativity_lat : test1.
|
||||
|
||||
Global Instance lattice_fun : Lattice (A -> B).
|
||||
Proof.
|
||||
split ; solve_fun.
|
||||
repeat (split; try (apply _)).
|
||||
eauto with test1.
|
||||
(* TODO *)
|
||||
all: solve_fun.
|
||||
apply associativity.
|
||||
apply commutativity.
|
||||
apply idempotency. apply _.
|
||||
apply associativity.
|
||||
apply commutativity.
|
||||
apply idempotency. apply _.
|
||||
Defined.
|
||||
|
||||
Global Instance boundedjoinsemilattice_fun
|
||||
`{@BoundedJoinSemiLattice B BJoin BBottom} :
|
||||
BoundedJoinSemiLattice (A -> B).
|
||||
Proof.
|
||||
repeat split; try apply _; try solve_fun.
|
||||
Defined.
|
||||
End fun_lattice.
|
||||
|
||||
@@ -92,24 +122,26 @@ End fun_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}.
|
||||
Context `{Bottom A}.
|
||||
Context {Hmax : forall x y, P x -> P y -> P (x ⊔ y)}.
|
||||
Context {Hmin : forall x y, P x -> P y -> P (x ⊓ y)}.
|
||||
Context {Hbot : P ⊥}.
|
||||
|
||||
Definition AP : Type := sig P.
|
||||
|
||||
Instance botAP : bottom AP := (empty_L ; Hbot).
|
||||
Instance botAP : Bottom AP.
|
||||
Proof. refine (⊥ ↾ _). apply Hbot. Defined.
|
||||
|
||||
Instance maxAP : maximum AP :=
|
||||
Instance maxAP : Join AP :=
|
||||
fun x y =>
|
||||
match x, y with
|
||||
| (a ; pa), (b ; pb) => (max_L0 a b ; Hmax a b pa pb)
|
||||
| (a ; pa), (b ; pb) => (a ⊔ b ; Hmax a b pa pb)
|
||||
end.
|
||||
|
||||
Instance minAP : minimum AP :=
|
||||
Instance minAP : Meet AP :=
|
||||
fun x y =>
|
||||
match x, y with
|
||||
| (a ; pa), (b ; pb) => (min_L0 a b ; Hmin a b pa pb)
|
||||
| (a ; pa), (b ; pb) => (a ⊓ b ; Hmin a b pa pb)
|
||||
end.
|
||||
|
||||
Instance hprop_sub : forall c, IsHProp (P c).
|
||||
@@ -127,19 +159,25 @@ Section sub_lattice.
|
||||
|
||||
Global Instance lattice_sub : Lattice AP.
|
||||
Proof.
|
||||
split ; solve_sub.
|
||||
repeat (split ; try (apply _ || solve_sub)).
|
||||
apply associativity.
|
||||
apply commutativity.
|
||||
apply idempotency. apply _.
|
||||
apply associativity.
|
||||
apply commutativity.
|
||||
apply idempotency. apply _.
|
||||
Defined.
|
||||
End sub_lattice.
|
||||
|
||||
Instance lor : maximum hProp := fun X Y => BuildhProp (Trunc (-1) (sum X Y)).
|
||||
Instance lor : Join 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.
|
||||
Instance land : Meet 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.
|
||||
@@ -181,17 +219,18 @@ Section hPropLattice.
|
||||
intros X Y Z.
|
||||
symmetry.
|
||||
apply path_hprop.
|
||||
symmetry.
|
||||
apply equiv_prod_assoc.
|
||||
Defined.
|
||||
|
||||
Instance lor_idempotent : Idempotent lor.
|
||||
Instance lor_idempotent : BinaryIdempotent lor.
|
||||
Proof.
|
||||
intros X.
|
||||
apply path_iff_hprop ; lor_intros
|
||||
; try(refine (tr(inl _))) ; auto.
|
||||
Defined.
|
||||
|
||||
Instance land_idempotent : Idempotent land.
|
||||
Instance land_idempotent : BinaryIdempotent land.
|
||||
Proof.
|
||||
intros X.
|
||||
apply path_iff_hprop ; cbn.
|
||||
@@ -199,14 +238,14 @@ Section hPropLattice.
|
||||
- intros a ; apply (pair a a).
|
||||
Defined.
|
||||
|
||||
Instance lor_neutrall : NeutralL lor lfalse.
|
||||
Instance lor_neutrall : LeftIdentity 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.
|
||||
Instance lor_neutralr : RightIdentity lor lfalse.
|
||||
Proof.
|
||||
intros X.
|
||||
apply path_iff_hprop ; lor_intros ; try contradiction
|
||||
@@ -235,16 +274,17 @@ Section hPropLattice.
|
||||
* 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 := _
|
||||
}.
|
||||
Global Instance lattice_hprop : Lattice hProp.
|
||||
Proof. repeat (split ; try apply _). Defined.
|
||||
|
||||
Global Instance bounded_jsl_hprop : BoundedJoinSemiLattice hProp.
|
||||
Proof. repeat (split ; try apply _). Qed.
|
||||
|
||||
Global Instance top_hprop : Top hProp := Unit_hp.
|
||||
Global Instance bounded_msl_hprop : @BoundedSemiLattice hProp (⊓) ⊤.
|
||||
Proof.
|
||||
repeat (split; try apply _); cbv.
|
||||
- intros X. apply path_trunctype ; apply prod_unit_l.
|
||||
- intros X. apply path_trunctype ; apply prod_unit_r.
|
||||
Defined.
|
||||
End hPropLattice.
|
||||
|
||||
@@ -1,119 +1,60 @@
|
||||
(** Interface for lattices and join semilattices. *)
|
||||
Require Import HoTT.
|
||||
From HoTT.Classes.interfaces Require Export abstract_algebra canonical_names.
|
||||
From HoTT.Classes Require Export theory.lattices.
|
||||
|
||||
(** Some preliminary notions to define lattices. *)
|
||||
Section binary_operation.
|
||||
Definition operation (A : Type) := A -> A -> A.
|
||||
|
||||
Variable (A : Type)
|
||||
(f : operation A).
|
||||
(* (** Join semilattices as a typeclass. They only have a join operator. *) *)
|
||||
(* Section JoinSemiLattice. *)
|
||||
(* Variable A : Type. *)
|
||||
(* Context {max_L : Join A} {empty_L : Bottom A}. *)
|
||||
|
||||
Class Commutative :=
|
||||
commutativity : forall x y, f x y = f y x.
|
||||
(* Class JoinSemiLattice := *)
|
||||
(* { *)
|
||||
(* commutative_max_js :> Commutative max_L ; *)
|
||||
(* associative_max_js :> Associative max_L ; *)
|
||||
(* idempotent_max_js :> BinaryIdempotent max_L ; *)
|
||||
(* neutralL_max_js :> LeftIdentity max_L empty_L ; *)
|
||||
(* neutralR_max_js :> RightIdentity max_L empty_L ; *)
|
||||
(* }. *)
|
||||
(* End JoinSemiLattice. *)
|
||||
|
||||
Class Associative :=
|
||||
associativity : forall x y z, f (f x y) z = f x (f y z).
|
||||
(* Arguments JoinSemiLattice _ {_} {_}. *)
|
||||
|
||||
Class Idempotent :=
|
||||
idempotency : forall x, f x x = x.
|
||||
(* Create HintDb joinsemilattice_hints. *)
|
||||
(* Hint Resolve associativity : joinsemilattice_hints. *)
|
||||
(* Hint Resolve (associativity _ _ _)^ : joinsemilattice_hints. *)
|
||||
(* Hint Resolve commutativity : joinsemilattice_hints. *)
|
||||
(* Hint Resolve idempotency : joinsemilattice_hints. *)
|
||||
(* Hint Resolve neutralityL : joinsemilattice_hints. *)
|
||||
(* Hint Resolve neutralityR : joinsemilattice_hints. *)
|
||||
|
||||
Variable g : operation A.
|
||||
(* (** Lattices as a typeclass which have both a join and a meet. *) *)
|
||||
(* Section Lattice. *)
|
||||
(* Variable A : Type. *)
|
||||
(* Context {max_L : maximum A} {min_L : minimum A} {empty_L : bottom A}. *)
|
||||
|
||||
Class Absorption :=
|
||||
absorb : forall x y, f x (g x y) = x.
|
||||
(* 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. *)
|
||||
|
||||
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 commutativity {_} {_} {_} _ _.
|
||||
Arguments associativity {_} {_} {_} _ _ _.
|
||||
Arguments idempotency {_} {_} {_} _.
|
||||
Arguments neutralityL {_} {_} {_} {_} _.
|
||||
Arguments neutralityR {_} {_} {_} {_} _.
|
||||
Arguments absorb {_} {_} {_} {_} _ _.
|
||||
|
||||
(** The operations in a lattice. *)
|
||||
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 {_}.
|
||||
|
||||
(** Join semilattices as a typeclass. They only have a join operator. *)
|
||||
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 commutativity : joinsemilattice_hints.
|
||||
Hint Resolve idempotency : joinsemilattice_hints.
|
||||
Hint Resolve neutralityL : joinsemilattice_hints.
|
||||
Hint Resolve neutralityR : joinsemilattice_hints.
|
||||
|
||||
(** Lattices as a typeclass which have both a join and a meet. *)
|
||||
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 _ {_} {_} {_}.
|
||||
(* Arguments Lattice _ {_} {_} {_}. *)
|
||||
|
||||
Create HintDb lattice_hints.
|
||||
Hint Resolve associativity : lattice_hints.
|
||||
Hint Resolve (associativity _ _ _)^ : lattice_hints.
|
||||
(* Hint Resolve (associativity _ _ _)^ : lattice_hints. *)
|
||||
Hint Resolve commutativity : lattice_hints.
|
||||
Hint Resolve absorb : lattice_hints.
|
||||
Hint Resolve absorption : lattice_hints.
|
||||
Hint Resolve idempotency : lattice_hints.
|
||||
Hint Resolve neutralityL : lattice_hints.
|
||||
Hint Resolve neutralityR : lattice_hints.
|
||||
Hint Resolve left_identity : lattice_hints.
|
||||
Hint Resolve right_identity : lattice_hints.
|
||||
|
||||
@@ -1,5 +1,7 @@
|
||||
(** The structure of a monad. *)
|
||||
(* TODO REMOVE THIS *)
|
||||
Require Import HoTT.
|
||||
Require Export HoTT.Classes.interfaces.monad.
|
||||
|
||||
Section functor.
|
||||
Variable (F : Type -> Type).
|
||||
@@ -11,35 +13,9 @@ Section functor.
|
||||
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
|
||||
}.
|
||||
Section monad_join.
|
||||
Context `{Monad M}.
|
||||
|
||||
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.
|
||||
Definition mjoin {A} (m : M (M A)) : M A := bind m id.
|
||||
End monad_join.
|
||||
|
||||
@@ -16,7 +16,7 @@ Section interface.
|
||||
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_filter : forall A φ X, f A (filter φ X) = {| f A X | φ |};
|
||||
f_member : forall A a X, member a X = a ∈ (f A X)
|
||||
}.
|
||||
|
||||
@@ -175,12 +175,12 @@ Section quotient_properties.
|
||||
Proof.
|
||||
intros ϕ X.
|
||||
apply (quotient_iso _)^-1.
|
||||
simple refine ({|_ & ϕ|}).
|
||||
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 & ϕ|}.
|
||||
{|class_of _ X | ϕ|} = class_of _ {|X | ϕ|}.
|
||||
Proof.
|
||||
rewrite <- (eissect (quotient_iso _)).
|
||||
simpl.
|
||||
@@ -224,30 +224,35 @@ Section quotient_properties.
|
||||
apply (ap _ HXY).
|
||||
Defined.
|
||||
|
||||
Instance View_max A : maximum (View A).
|
||||
Instance join_view A : Join (View A).
|
||||
Proof.
|
||||
apply view_union.
|
||||
Defined.
|
||||
|
||||
Hint Unfold Commutative Associative Idempotent NeutralL NeutralR View_max view_union.
|
||||
Local Hint Unfold Commutative Associative HeteroAssociative Idempotent BinaryIdempotent LeftIdentity RightIdentity join_view view_union sg_op join_is_sg_op meet_is_sg_op.
|
||||
|
||||
Instance bottom_view A : bottom (View A).
|
||||
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 sq1 := autounfold; intros;
|
||||
unfold view_union; try f_ap;
|
||||
rewrite ?(eisretr (quotient_iso _)).
|
||||
|
||||
Ltac sq2 := autounfold; intros;
|
||||
unfold view_union;
|
||||
rewrite <- (eissect (quotient_iso _)), ?(eisretr (quotient_iso _));
|
||||
f_ap; simpl.
|
||||
|
||||
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).
|
||||
Global Instance view_lattice A : BoundedJoinSemiLattice (View A).
|
||||
Proof.
|
||||
split ; try sq1 ; try sq2.
|
||||
repeat split; try apply _.
|
||||
- sq1. apply associativity.
|
||||
- sq2. apply left_identity.
|
||||
- sq2. apply right_identity.
|
||||
- sq1. apply commutativity.
|
||||
- sq2. apply binary_idempotent.
|
||||
Defined.
|
||||
|
||||
End quotient_properties.
|
||||
@@ -316,7 +321,8 @@ Section properties.
|
||||
|
||||
Ltac via_quotient := intros ; apply reflect_f_eq ; simpl
|
||||
; rewrite <- ?(well_defined_union T _), <- ?(well_defined_empty T _)
|
||||
; eauto with lattice_hints typeclass_instances.
|
||||
; try (apply commutativity || apply associativity || apply binary_idempotent || apply left_identity || apply right_identity).
|
||||
(* TODO ; eauto with lattice_hints typeclass_instances. *)
|
||||
|
||||
Lemma union_comm : forall A (X Y : T A),
|
||||
set_eq f (X ∪ Y) (Y ∪ X).
|
||||
@@ -328,6 +334,7 @@ Section properties.
|
||||
set_eq f ((X ∪ Y) ∪ Z) (X ∪ (Y ∪ Z)).
|
||||
Proof.
|
||||
via_quotient.
|
||||
symmetry. apply associativity.
|
||||
Defined.
|
||||
|
||||
Lemma union_idem : forall A (X : T A),
|
||||
@@ -376,4 +383,4 @@ Section refinement.
|
||||
(h : FSet A -> B)
|
||||
: T A -> B
|
||||
:= fun X => h(quotient_iso (f A) (class_of _ X)).
|
||||
End refinement.
|
||||
End refinement.
|
||||
|
||||
@@ -53,8 +53,8 @@ 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).
|
||||
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).
|
||||
Infix "⊆_d" := subset_dec (at level 10, right associativity).
|
||||
|
||||
Reference in New Issue
Block a user