Skip to content

Commit

Permalink
refactor into Base, Properties with re-export from Refinement
Browse files Browse the repository at this point in the history
  • Loading branch information
jamesmckinna committed Oct 7, 2024
1 parent 65d9eac commit 0b8f459
Show file tree
Hide file tree
Showing 3 changed files with 69 additions and 34 deletions.
9 changes: 8 additions & 1 deletion CHANGELOG.md
Original file line number Diff line number Diff line change
Expand Up @@ -110,8 +110,9 @@ New modules

* `Data.List.Relation.Binary.Permutation.Propositional.Properties.WithK`

* Properties of `Data.Refinement` types:
* Refactored `Data.Refinement` into:
```agda
Data.Refinement.Base
Data.Refinement.Properties
```

Expand Down Expand Up @@ -317,6 +318,12 @@ Additions to existing modules
m≤pred[n]⇒suc[m]≤n : .{{NonZero n}} → m ≤ pred n → suc m ≤ n
```

* New properties re-exported from `Data.Refinement`:
```agda
value-injective : value v ≡ value w → v ≡ w
_≟_ : DecidableEquality A → DecidableEquality [ x ∈ A ∣ P x ]
```

* New lemma in `Data.Vec.Properties`:
```agda
map-concat : map f (concat xss) ≡ concat (map (map f) xss)
Expand Down
40 changes: 7 additions & 33 deletions src/Data/Refinement.agda
Original file line number Diff line number Diff line change
Expand Up @@ -8,39 +8,13 @@

module Data.Refinement where

open import Level
open import Data.Irrelevant as Irrelevant using (Irrelevant)
open import Function.Base
open import Relation.Unary using (IUniversal; _⇒_; _⊢_)

private
variable
a b p q : Level
A : Set a
B : Set b

record Refinement {a p} (A : Set a) (P : A Set p) : Set (a ⊔ p) where
constructor _,_
field value : A
proof : Irrelevant (P value)
infixr 4 _,_
open Refinement public

-- The syntax declaration below is meant to mimick set comprehension.
-- It is attached to Refinement-syntax, to make it easy to import
-- Data.Refinement without the special syntax.

infix 2 Refinement-syntax
Refinement-syntax = Refinement
syntax Refinement-syntax A (λ x P) = [ x ∈ A ∣ P ]

module _ {P : A Set p} {Q : B Set q} where
------------------------------------------------------------------------
-- Publicly re-export the contents of the base module

map : (f : A B) ∀[ P ⇒ f ⊢ Q ]
[ a ∈ A ∣ P a ] [ b ∈ B ∣ Q b ]
map f prf (a , p) = f a , Irrelevant.map prf p
open import Data.Refinement.Base public

module _ {P : A Set p} {Q : A Set q} where
------------------------------------------------------------------------
-- Publicly re-export queries

refine : ∀[ P ⇒ Q ] [ a ∈ A ∣ P a ] [ a ∈ A ∣ Q a ]
refine = map id
open import Data.Refinement.Properties public
using (value-injective; _≟_)
54 changes: 54 additions & 0 deletions src/Data/Refinement/Base.agda
Original file line number Diff line number Diff line change
@@ -0,0 +1,54 @@
------------------------------------------------------------------------
-- The Agda standard library
--
-- Refinement type: a value together with a proof irrelevant witness.
------------------------------------------------------------------------

{-# OPTIONS --cubical-compatible --safe #-}

module Data.Refinement.Base where

open import Level
open import Data.Irrelevant as Irrelevant using (Irrelevant)
open import Function.Base
open import Relation.Unary using (IUniversal; _⇒_; _⊢_)

private
variable
a b p q : Level
A : Set a
B : Set b
P : A Set p


------------------------------------------------------------------------
-- Definition

record Refinement (A : Set a) (P : A Set p) : Set (a ⊔ p) where
constructor _,_
field value : A
proof : Irrelevant (P value)
infixr 4 _,_
open Refinement public

-- The syntax declaration below is meant to mimic set comprehension.
-- It is attached to Refinement-syntax, to make it easy to import
-- Data.Refinement without the special syntax.

infix 2 Refinement-syntax
Refinement-syntax = Refinement
syntax Refinement-syntax A (λ x P) = [ x ∈ A ∣ P ]

------------------------------------------------------------------------
-- Basic operations

module _ {Q : B Set q} where

map : (f : A B) ∀[ P ⇒ f ⊢ Q ]
[ a ∈ A ∣ P a ] [ b ∈ B ∣ Q b ]
map f prf (a , p) = f a , Irrelevant.map prf p

module _ {Q : A Set q} where

refine : ∀[ P ⇒ Q ] [ a ∈ A ∣ P a ] [ a ∈ A ∣ Q a ]
refine = map id

0 comments on commit 0b8f459

Please sign in to comment.