102 lines
2.7 KiB
Agda
102 lines
2.7 KiB
Agda
{-# OPTIONS --safe --without-K #-}
|
||
|
||
open import Data.Nat.Base hiding (_+_)
|
||
|
||
module Regex (α : ℕ) where
|
||
|
||
import Algebra.Definitions
|
||
open import Algebra.Structures
|
||
open import Data.Bool.Base
|
||
open import Data.Fin.Base using (Fin)
|
||
open import Data.Fin.Properties using (_≟_)
|
||
open import Data.List.Base
|
||
open import Data.Product.Base
|
||
open import Data.Sum.Base
|
||
open import Function.Base
|
||
open import Relation.Nullary
|
||
open import Relation.Unary
|
||
open import Relation.Binary hiding (Decidable)
|
||
|
||
-- Note: we do not define negation or intersections.
|
||
-- These can be derived, but only by a double-exponential blow-up in space.
|
||
-- So in the future we should consider adding them.
|
||
data Regex : Set where
|
||
-- Do not recognize any strings
|
||
Λ : Regex
|
||
-- Union
|
||
_+_ : Regex → Regex → Regex
|
||
-- Sequencing: first one, and then the other
|
||
_⨟_ : Regex → Regex → Regex
|
||
-- Kleene star
|
||
_⋆ : Regex → Regex
|
||
-- Recognize a single character
|
||
⁅_⁆ : Fin α → Regex
|
||
|
||
ε : Regex
|
||
ε = Λ ⋆
|
||
|
||
-- Do we accept the empty string?
|
||
data A : Regex → Set where
|
||
+ₗ : ∀ {r₁ r₂} → A r₁ → A (r₁ + r₂)
|
||
+ᵣ : ∀ {r₁ r₂} → A r₂ → A (r₁ + r₂)
|
||
_⨟_ : ∀ {r₁ r₂} → A r₁ → A r₂ → A (r₁ ⨟ r₂)
|
||
star : ∀ {r} → A (r ⋆)
|
||
|
||
-- ε accepts the empty string
|
||
Aε : A ε
|
||
Aε = star
|
||
|
||
-- A is decidable!
|
||
A? : Decidable A
|
||
A? Λ = no λ ()
|
||
A? (r₁ + r₂) = map′
|
||
[ +ₗ , +ᵣ ]′
|
||
(λ { (+ₗ Ar₁) → inj₁ Ar₁
|
||
; (+ᵣ Ar₂) → inj₂ Ar₂
|
||
}
|
||
)
|
||
(A? r₁ ⊎-dec A? r₂)
|
||
A? (r₁ ⨟ r₂) = map′
|
||
(uncurry _⨟_)
|
||
(λ { (Ar₁ ⨟ Ar₂) → Ar₁ , Ar₂})
|
||
(A? r₁ ×-dec A? r₂)
|
||
A? (r ⋆) = yes star
|
||
A? ⁅ x ⁆ = no λ ()
|
||
|
||
-- Take the Brzozowski derivative of a regular expression by a
|
||
-- character
|
||
δ : Regex → Fin α → Regex
|
||
δ Λ i = Λ
|
||
δ (r₁ + r₂) i = δ r₁ i + δ r₂ i
|
||
δ (r₁ ⨟ r₂) i = if does (A? r₁)
|
||
then δ r₂ i + (δ r₁ i ⨟ r₂)
|
||
else (δ r₁ i ⨟ r₂)
|
||
δ (r ⋆) i = δ r i ⨟ (r ⋆)
|
||
δ ⁅ x ⁆ i = if does (x ≟ i) then ε else Λ
|
||
|
||
-- Take the Brzozowski derivative of a regular expression by a
|
||
-- string of characters
|
||
δ⋆ : Regex → List (Fin α) → Regex
|
||
δ⋆ = foldl δ
|
||
|
||
-- Whether a regular expression expresses a particular string
|
||
Accepts : Regex → List (Fin α) → Set
|
||
Accepts r as = A (δ⋆ r as)
|
||
|
||
-- Accepts is decidable!
|
||
Accepts? : ∀ r as → Dec (Accepts r as)
|
||
Accepts? r as = A? (δ⋆ r as)
|
||
|
||
------------------------------------------------------------------------
|
||
-- Relations on regeces
|
||
|
||
_⊑_ : Rel Regex _
|
||
r₁ ⊑ r₂ = ∀ as → Accepts r₁ as → Accepts r₂ as
|
||
|
||
_⊒_ : Rel Regex _
|
||
r₁ ⊒ r₂ = r₂ ⊑ r₁
|
||
|
||
_≈_ : Rel Regex _
|
||
r₁ ≈ r₂ = (r₁ ⊑ r₂) × (r₁ ⊒ r₂)
|
||
|