Skip to content

Commit d81855d

Browse files
authored
Monotonicity and modalities for predicates over types with orders (#10)
* (new) monotonicity and modalities for predicates over types with a partial order * Add to index, solve name clash
1 parent 73f4da0 commit d81855d

File tree

3 files changed

+91
-0
lines changed

3 files changed

+91
-0
lines changed

Class/MonotonePredicate.agda

+3
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,3 @@
1+
module Class.MonotonePredicate where
2+
3+
open import Class.MonotonePredicate.Core public

Class/MonotonePredicate/Core.agda

+87
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,87 @@
1+
2+
module Class.MonotonePredicate.Core where
3+
4+
open import Class.Prelude
5+
open import Class.HasOrder
6+
7+
open import Relation.Unary
8+
open import Relation.Binary using (IsPreorder ; IsEquivalence)
9+
10+
open import Level
11+
12+
-- We make the "simplifying" assumption here that everything (carrier,
13+
-- relation, etc...) lives at the same level. Otherwise we couldn't
14+
-- express the (co)monadic structure of the possiblity and necessity
15+
-- modalities.
16+
module _ {A : Type ℓ}
17+
⦃ _ : HasPreorder {A = A} {_≈_ = _≡_} {ℓ} {ℓ} ⦄ where
18+
19+
open IsPreorder {ℓ} ≤-isPreorder
20+
21+
record Monotone {ℓ′} (P : Pred A ℓ′) : Type (ℓ ⊔ ℓ′ ⊔ ℓ) where
22+
field
23+
weaken : {a a′} a ≤ a′ P a P a′
24+
25+
record Antitone {ℓ′} (P : Pred A ℓ′) : Type (ℓ ⊔ ℓ′ ⊔ ℓ″) where
26+
field
27+
strengthen : {a a′} a′ ≤ a P a P a′
28+
29+
open Monotone ⦃...⦄
30+
open Antitone ⦃...⦄
31+
32+
-- The posibility modality. One way to think about posibility is as a unary
33+
-- version of separating conjunction.
34+
record (P : Pred A ℓ) (a : A) : Setwhere
35+
constructor ◇⟨_,_⟩
36+
field
37+
{a′} : A
38+
ι : a′ ≤ a
39+
px : P a′
40+
41+
open public
42+
43+
-- The necessity modality. In a similar spirit, we can view necessity as a unary
44+
-- version of separating implication.
45+
record (P : Pred A ℓ) (a : A) : Setwhere
46+
constructor necessary
47+
field
48+
□⟨_⟩_ : {a′} : a ≤ a′) P a′
49+
50+
open public
51+
52+
module Operations where
53+
54+
-- □ is a comonad over the category of monotone predicates over `A`
55+
extract : {P} ∀[ □ P ⇒ P ]
56+
extract px = □⟨ px ⟩ reflexive _≡_.refl
57+
58+
duplicate : {P} ∀[ □ P ⇒ □ (□ P) ]
59+
duplicate px = necessary λ ι necessary λ ι′ □⟨ px ⟩ trans ι ι′
60+
61+
62+
-- ◇ is a monad over the category of monotone predicates over `A`.
63+
return : {P} ∀[ P ⇒ ◇ P ]
64+
return px = ◇⟨ reflexive _≡_.refl , px ⟩
65+
66+
join : {P} ∀[ ◇ (◇ P) ⇒ ◇ P ]
67+
join ◇⟨ ι₁ , ◇⟨ ι₂ , px ⟩ ⟩ = ◇⟨ (trans ι₂ ι₁) , px ⟩
68+
69+
-- □ is right-adjoint to ◇
70+
curry : {P Q} ∀[ ◇ P ⇒ Q ] ∀[ P ⇒ □ Q ]
71+
curry f px = necessary (λ ι f ◇⟨ ι , px ⟩)
72+
73+
uncurry : {P Q} ∀[ P ⇒ □ Q ] ∀[ ◇ P ⇒ Q ]
74+
uncurry f ◇⟨ ι , px ⟩ = □⟨ f px ⟩ ι
75+
76+
77+
-- The "Kripke exponent" (or, Kripke function space) between two predicates is
78+
-- defined as the necessity of their implication.
79+
_⇛_ : (P Q : Pred A ℓ) Pred A ℓ
80+
P ⇛ Q = □ (P ⇒ Q)
81+
82+
kripke-curry : {P Q R : Pred A ℓ} ⦃ Monotone P ⦄ ∀[ (P ∩ Q) ⇒ R ] ∀[ P ⇒ (Q ⇛ R) ]
83+
kripke-curry f px₁ = necessary (λ i px₂ f (weaken i px₁ , px₂))
84+
85+
kripke-uncurry : {P Q R : Pred A ℓ} ∀[ P ⇒ (Q ⇛ R) ] ∀[ (P ∩ Q) ⇒ R ]
86+
kripke-uncurry f (px₁ , px₂) = □⟨ f px₁ ⟩ reflexive _≡_.refl $ px₂
87+

standard-library-classes.agda

+1
Original file line numberDiff line numberDiff line change
@@ -25,6 +25,7 @@ open import Class.HasAdd public
2525
open import Class.HasOrder public
2626
open import Class.Show public
2727
open import Class.ToBool public
28+
open import Class.MonotonePredicate public
2829

2930
-- ** Tests
3031
open import Test.Monoid

0 commit comments

Comments
 (0)