2017-08-01 15:12:59 +02:00
|
|
|
(* Typeclass for lattices *)
|
2017-06-19 17:08:56 +02:00
|
|
|
Require Import HoTT.
|
|
|
|
|
2017-08-03 12:21:34 +02:00
|
|
|
Section binary_operation.
|
|
|
|
Variable A : Type.
|
|
|
|
|
|
|
|
Definition operation := A -> A -> A.
|
|
|
|
|
|
|
|
Class maximum :=
|
|
|
|
max_L : operation.
|
|
|
|
|
|
|
|
Class minimum :=
|
|
|
|
min_L : operation.
|
|
|
|
|
|
|
|
Class bottom :=
|
|
|
|
empty : A.
|
|
|
|
End binary_operation.
|
|
|
|
|
2017-08-03 17:09:10 +02:00
|
|
|
Arguments max_L {_} {_} _.
|
|
|
|
Arguments min_L {_} {_} _.
|
2017-08-03 12:21:34 +02:00
|
|
|
Arguments empty {_}.
|
2017-06-19 17:08:56 +02:00
|
|
|
|
|
|
|
Section Defs.
|
|
|
|
Variable A : Type.
|
|
|
|
Variable f : A -> A -> 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 :=
|
2017-08-03 12:21:34 +02:00
|
|
|
absorb : forall x y, f x (g x y) = x.
|
2017-06-19 17:08:56 +02:00
|
|
|
|
|
|
|
Variable n : A.
|
|
|
|
|
|
|
|
Class NeutralL :=
|
2017-08-01 17:25:57 +02:00
|
|
|
neutralityL : forall x, f n x = x.
|
2017-06-19 17:08:56 +02:00
|
|
|
|
|
|
|
Class NeutralR :=
|
2017-08-01 17:25:57 +02:00
|
|
|
neutralityR : forall x, f x n = x.
|
2017-06-19 17:08:56 +02:00
|
|
|
|
|
|
|
End Defs.
|
|
|
|
|
|
|
|
Arguments Commutative {_} _.
|
|
|
|
Arguments Associative {_} _.
|
|
|
|
Arguments Idempotent {_} _.
|
|
|
|
Arguments NeutralL {_} _ _.
|
|
|
|
Arguments NeutralR {_} _ _.
|
|
|
|
Arguments Absorption {_} _ _.
|
|
|
|
|
2017-08-08 13:45:27 +02:00
|
|
|
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 commutative : joinsemilattice_hints.
|
|
|
|
Hint Resolve idempotency : joinsemilattice_hints.
|
|
|
|
Hint Resolve neutralityL : joinsemilattice_hints.
|
|
|
|
Hint Resolve neutralityR : joinsemilattice_hints.
|
|
|
|
|
2017-06-19 17:08:56 +02:00
|
|
|
Section Lattice.
|
|
|
|
Variable A : Type.
|
2017-08-03 12:21:34 +02:00
|
|
|
Context {max_L : maximum A} {min_L : minimum A} {empty_L : bottom A}.
|
2017-06-19 17:08:56 +02:00
|
|
|
|
|
|
|
Class Lattice :=
|
|
|
|
{
|
2017-08-03 12:21:34 +02:00
|
|
|
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
|
2017-06-19 17:08:56 +02:00
|
|
|
}.
|
|
|
|
End Lattice.
|
|
|
|
|
2017-08-03 12:21:34 +02:00
|
|
|
Arguments Lattice _ {_} {_} {_}.
|
2017-06-19 17:08:56 +02:00
|
|
|
|
2017-08-03 12:21:34 +02:00
|
|
|
Create HintDb 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.
|
2017-06-19 17:08:56 +02:00
|
|
|
|
|
|
|
Section BoolLattice.
|
|
|
|
|
2017-08-02 14:21:12 +02:00
|
|
|
Ltac solve_bool :=
|
2017-06-19 21:32:55 +02:00
|
|
|
let x := fresh in
|
|
|
|
repeat (intro x ; destruct x)
|
|
|
|
; compute
|
|
|
|
; auto
|
|
|
|
; try contradiction.
|
|
|
|
|
2017-08-03 12:21:34 +02:00
|
|
|
Instance maximum_bool : maximum Bool := orb.
|
|
|
|
Instance minimum_bool : minimum Bool := andb.
|
|
|
|
Instance bottom_bool : bottom Bool := false.
|
|
|
|
|
|
|
|
Global Instance lattice_bool : Lattice Bool.
|
2017-06-19 17:08:56 +02:00
|
|
|
Proof.
|
2017-08-03 12:21:34 +02:00
|
|
|
split ; solve_bool.
|
2017-06-19 17:08:56 +02:00
|
|
|
Defined.
|
|
|
|
|
2017-08-01 15:12:59 +02:00
|
|
|
Definition and_true : forall b, andb b true = b.
|
|
|
|
Proof.
|
2017-08-02 14:21:12 +02:00
|
|
|
solve_bool.
|
2017-08-01 15:12:59 +02:00
|
|
|
Defined.
|
|
|
|
|
|
|
|
Definition and_false : forall b, andb b false = false.
|
|
|
|
Proof.
|
2017-08-02 14:21:12 +02:00
|
|
|
solve_bool.
|
2017-08-01 15:12:59 +02:00
|
|
|
Defined.
|
|
|
|
|
|
|
|
Definition dist₁ : forall b₁ b₂ b₃,
|
|
|
|
andb b₁ (orb b₂ b₃) = orb (andb b₁ b₂) (andb b₁ b₃).
|
|
|
|
Proof.
|
2017-08-02 14:21:12 +02:00
|
|
|
solve_bool.
|
2017-08-01 15:12:59 +02:00
|
|
|
Defined.
|
|
|
|
|
|
|
|
Definition dist₂ : forall b₁ b₂ b₃,
|
|
|
|
orb b₁ (andb b₂ b₃) = andb (orb b₁ b₂) (orb b₁ b₃).
|
|
|
|
Proof.
|
2017-08-02 14:21:12 +02:00
|
|
|
solve_bool.
|
2017-08-01 15:12:59 +02:00
|
|
|
Defined.
|
|
|
|
|
|
|
|
Definition max_min : forall b₁ b₂,
|
|
|
|
orb (andb b₁ b₂) b₁ = b₁.
|
|
|
|
Proof.
|
2017-08-02 14:21:12 +02:00
|
|
|
solve_bool.
|
2017-08-01 15:12:59 +02:00
|
|
|
Defined.
|
|
|
|
|
2017-06-19 17:08:56 +02:00
|
|
|
End BoolLattice.
|
|
|
|
|
2017-08-02 14:21:12 +02:00
|
|
|
Section fun_lattice.
|
2017-08-03 12:21:34 +02:00
|
|
|
Context {A B : Type}.
|
|
|
|
Context `{Lattice B}.
|
2017-08-02 14:21:12 +02:00
|
|
|
Context `{Funext}.
|
|
|
|
|
2017-08-03 12:21:34 +02:00
|
|
|
Global Instance max_fun : maximum (A -> B) :=
|
|
|
|
fun (f g : A -> B) (a : A) => max_L0 (f a) (g a).
|
2017-08-02 14:21:12 +02:00
|
|
|
|
2017-08-03 12:21:34 +02:00
|
|
|
Global Instance min_fun : minimum (A -> B) :=
|
|
|
|
fun (f g : A -> B) (a : A) => min_L0 (f a) (g a).
|
2017-08-02 14:21:12 +02:00
|
|
|
|
2017-08-03 12:21:34 +02:00
|
|
|
Global Instance bot_fun : bottom (A -> B)
|
|
|
|
:= fun _ => empty_L.
|
2017-08-02 14:21:12 +02:00
|
|
|
|
2017-08-03 12:21:34 +02:00
|
|
|
Ltac solve_fun :=
|
|
|
|
compute ; intros ; apply path_forall ; intro ;
|
|
|
|
eauto with lattice_hints typeclass_instances.
|
2017-08-02 14:21:12 +02:00
|
|
|
|
2017-08-03 12:21:34 +02:00
|
|
|
Global Instance lattice_fun : Lattice (A -> B).
|
2017-08-02 14:21:12 +02:00
|
|
|
Proof.
|
2017-08-03 12:21:34 +02:00
|
|
|
split ; solve_fun.
|
2017-08-02 14:21:12 +02:00
|
|
|
Defined.
|
|
|
|
|
|
|
|
End fun_lattice.
|
|
|
|
|
|
|
|
Section sub_lattice.
|
2017-08-03 12:21:34 +02:00
|
|
|
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}.
|
2017-08-02 14:21:12 +02:00
|
|
|
|
|
|
|
Definition AP : Type := sig P.
|
|
|
|
|
2017-08-03 12:21:34 +02:00
|
|
|
Instance botAP : bottom AP := (empty_L ; Hbot).
|
2017-08-02 14:21:12 +02:00
|
|
|
|
2017-08-03 12:21:34 +02:00
|
|
|
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.
|
2017-08-02 14:21:12 +02:00
|
|
|
|
2017-08-03 12:21:34 +02:00
|
|
|
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.
|
2017-08-02 14:21:12 +02:00
|
|
|
|
|
|
|
Instance hprop_sub : forall c, IsHProp (P c).
|
|
|
|
Proof.
|
|
|
|
apply _.
|
|
|
|
Defined.
|
|
|
|
|
|
|
|
Ltac solve_sub :=
|
|
|
|
let x := fresh in
|
|
|
|
repeat (intro x ; destruct x)
|
2017-08-03 12:21:34 +02:00
|
|
|
; simple refine (path_sigma _ _ _ _ _)
|
|
|
|
; simpl
|
|
|
|
; try (apply hprop_sub)
|
|
|
|
; eauto 3 with lattice_hints typeclass_instances.
|
2017-08-02 14:21:12 +02:00
|
|
|
|
2017-08-03 12:21:34 +02:00
|
|
|
Global Instance lattice_sub : Lattice AP.
|
2017-08-02 14:21:12 +02:00
|
|
|
Proof.
|
2017-08-03 12:21:34 +02:00
|
|
|
split ; solve_sub.
|
2017-08-02 14:21:12 +02:00
|
|
|
Defined.
|
|
|
|
|
|
|
|
End sub_lattice.
|
|
|
|
|
2017-08-03 12:21:34 +02:00
|
|
|
Create HintDb 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.
|
|
|
|
|
2017-06-19 17:54:44 +02:00
|
|
|
Hint Resolve
|
2017-08-03 12:21:34 +02:00
|
|
|
associativity
|
|
|
|
and_true and_false
|
2017-08-01 15:12:59 +02:00
|
|
|
dist₁ dist₂ max_min
|
2017-06-19 17:54:44 +02:00
|
|
|
: bool_lattice_hints.
|