├── .gitignore ├── report ├── report.pdf ├── latex │ ├── logo.jpg │ ├── logohd.png │ ├── bibliography.tex │ ├── first_page.tex │ ├── conclusion.tex │ ├── main.tex │ ├── my-lexer.py │ ├── agda_introduction.tex │ └── regular_expressions.tex └── introduction_examples.agda ├── README.md ├── Regexp ├── Equivalence.agda ├── String.agda ├── ENfa.agda ├── VecUtil.agda ├── DfaNfa.agda ├── SubsetFinIso.agda ├── Brzozowski.agda ├── RegexpNfa.agda ├── RegExpSet.agda ├── examples.agda ├── Dfa.agda ├── LTS.agda ├── Regexp.agda └── PumpingLemmaExample.agda └── plfa ├── equality.agda ├── negation.agda ├── decidable.agda ├── isomorphism.agda ├── introduction.agda ├── connectives.agda ├── quantifiers.agda ├── relations.agda └── lists.agda /.gitignore: -------------------------------------------------------------------------------- 1 | *.agdai 2 | MAlonzo 3 | -------------------------------------------------------------------------------- /report/report.pdf: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/desi-ivanov/agda-regexp-automata/HEAD/report/report.pdf -------------------------------------------------------------------------------- /report/latex/logo.jpg: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/desi-ivanov/agda-regexp-automata/HEAD/report/latex/logo.jpg -------------------------------------------------------------------------------- /report/latex/logohd.png: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/desi-ivanov/agda-regexp-automata/HEAD/report/latex/logohd.png -------------------------------------------------------------------------------- /README.md: -------------------------------------------------------------------------------- 1 | # Formalization of Regular Languages in Agda 2 | We use the proof assistant Agda to prove properties on Regular Expressions and Finite State Automata: 3 | - [Regular Expressions and algebraic laws](Regexp/Regexp.agda) 4 | - [The matching problem solved using Brozozowski's Derivatives](Regexp/Brzozowski.agda) 5 | - [Deterministic](Regexp/Dfa.agda) and [non-deterministic](Regexp/Nfa.agda) finite state automata 6 | - [The matching problem solved using finite state automata](Regexp/RegexpNfa.agda) 7 | - [Pumping lemma for regular languages](Regexp/Dfa.agda) 8 | 9 | A final report is available [here](report/report.pdf) 10 | -------------------------------------------------------------------------------- /report/latex/bibliography.tex: -------------------------------------------------------------------------------- 1 | \begin{thebibliography}{9} 2 | 3 | \bibitem{PLFA} 4 | Wadler, Philip and Wen Kokke. Programming Language Foundations in Agda. Available at \texttt{http://plfa.inf.ed.ac.uk}. 2019. 5 | 6 | \bibitem{AGDA WIKI} 7 | Agda Wiki. Chalmers and Gothenburg University, 2.4 edn. (2014), \texttt{http://wiki.portal.chalmers.se/agda} 8 | 9 | \bibitem{LFT} 10 | Hopcroft, John E.; Motwani, Rajeev; Ullman, Jeffrey D. (2013). Introduction to Automata Theory, Languages, and Computation (3rd ed.). 11 | 12 | \bibitem{Brzozowski64derivativesof} 13 | Janusz A. Brzozowski, Derivatives of regular expressions, JOURNAL OF THE ACM, 1964, 11, 481--494 14 | 15 | \end{thebibliography} -------------------------------------------------------------------------------- /report/latex/first_page.tex: -------------------------------------------------------------------------------- 1 | \newgeometry{centering} 2 | 3 | \begin{titlepage} 4 | 5 | \centerline {\Large{\textsc{ UNIVERSIT\`A DEGLI STUDI DI TORINO}}} 6 | \vskip 20 pt 7 | 8 | \centerline {\Large{\textsc DIPARTIMENTO DI INFORMATICA}} 9 | 10 | \vskip 20 pt 11 | 12 | \centerline {{\textsc SCUOLA DI SCIENZE DELLA NATURA}} 13 | 14 | \vskip 20 pt 15 | 16 | \centerline {\Large{\textsc Laurea Triennale in Informatica}} 17 | 18 | 19 | \vskip 60 pt 20 | 21 | 22 | 23 | %\begin{tabular}{ccc} 24 | \centerline {\includegraphics[width=5cm]{logohd.png}} 25 | % \end{tabular} 26 | 27 | \vskip 1.2cm 28 | 29 | \vskip 0.7cm 30 | 31 | \centerline {\Large {\bf Formalization of Regular Languages in Agda}} 32 | 33 | \vskip 1.7cm 34 | 35 | 36 | \noindent Relatore: Prof. Luca Padovani 37 | \hfill {Candidato: Desislav Nikolaev Ivanov} 38 | 39 | 40 | 41 | \vskip 2.3cm 42 | \centerline{ANNO ACCADEMICO} 43 | \centerline{2019 / 2020} 44 | \end{titlepage} 45 | % \thispagestyle{empty} 46 | \restoregeometry -------------------------------------------------------------------------------- /Regexp/Equivalence.agda: -------------------------------------------------------------------------------- 1 | module Equivalence where 2 | open import Relation.Binary.PropositionalEquality using (_≡_) 3 | open import Data.Nat as ℕ using (ℕ; zero; suc; _≥_; _≤_; _<_; s≤s; z≤n; _+_; _*_; _^_) 4 | open import Data.Nat.Properties 5 | open import Data.Fin using (Fin; toℕ; inject+; raise; 0F; 1F; 2F; 3F; 4F; 5F; 6F) renaming (_<_ to _<ᶠ_; zero to fzero; suc to fsuc) 6 | open import Data.Fin.Properties using (pigeonhole; toℕ-inject+) 7 | open import Data.Product using (_×_; Σ; ∃; ∃₂; Σ-syntax; ∃-syntax; _,_; proj₁; proj₂) 8 | open import Data.Fin.Subset 9 | using (Subset; ⁅_⁆; _∪_; _∩_; _∈_; Nonempty) 10 | renaming (⊥ to ∅; ⊤ to FullSet) 11 | 12 | infix 0 _≃_ 13 | record _≃_ (A B : Set) : Set where 14 | field 15 | to : A → B 16 | from : B → A 17 | from∘to : ∀ (x : A) → from (to x) ≡ x 18 | to∘from : ∀ (y : B) → to (from y) ≡ y 19 | 20 | infix 0 _⇔_ 21 | record _⇔_ (A B : Set) : Set where 22 | field 23 | to : A → B 24 | from : B → A 25 | 26 | _IFF_ : {A B : Set} → (A → B) → (B → A) → A ⇔ B 27 | f IFF g = record { to = f ; from = g } 28 | 29 | 30 | infix 0 _≲_ 31 | record _≲_ (A B : Set) : Set where 32 | field 33 | to : A → B 34 | from : B → A 35 | from∘to : ∀ (x : A) → from (to x) ≡ x 36 | -------------------------------------------------------------------------------- /Regexp/String.agda: -------------------------------------------------------------------------------- 1 | open import Data.Nat using (ℕ; zero; suc) 2 | import Relation.Binary.PropositionalEquality as Eq 3 | open Eq using (_≡_; refl) 4 | open Eq.≡-Reasoning 5 | open Eq.≡-Reasoning using (begin_; _≡⟨⟩_; _≡⟨_⟩_; _∎) 6 | module String (Σ : Set) where 7 | 8 | infixr 5 _∷_ 9 | data String : Set where 10 | [] : String 11 | _∷_ : Σ → String → String 12 | 13 | 14 | infixr 5 _++_ 15 | _++_ : String → String → String 16 | [] ++ ys = ys 17 | (x ∷ xs) ++ ys = x ∷ (xs ++ ys) 18 | 19 | ++-assoc : ∀ (s t g : String) → (s ++ t) ++ g ≡ s ++ (t ++ g) 20 | ++-assoc [] t g = refl 21 | ++-assoc (x ∷ s) t g rewrite ++-assoc s t g = refl 22 | 23 | ++-idʳ : ∀ (s : String) → (s ++ []) ≡ s 24 | ++-idʳ [] = refl 25 | ++-idʳ (x ∷ s) rewrite ++-idʳ s = refl 26 | 27 | ++-idˡ : ∀ (s : String) → ([] ++ s) ≡ s 28 | ++-idˡ [] = refl 29 | ++-idˡ (x ∷ x₁) = refl 30 | 31 | foldl : ∀ {B : Set} → (B → Σ → B) → B → String → B 32 | foldl f b [] = b 33 | foldl f b (x ∷ xs) = foldl f (f b x) xs 34 | 35 | length : String → ℕ 36 | length [] = zero 37 | length (x ∷ xs) = suc (length xs) 38 | 39 | take : ℕ → String → String 40 | take zero xs = [] 41 | take (suc n) [] = [] 42 | take (suc n) (x ∷ xs) = x ∷ take n xs 43 | 44 | drop : ℕ → String → String 45 | drop zero xs = xs 46 | drop (suc n) [] = [] 47 | drop (suc n) (x ∷ xs) = drop n xs 48 | -------------------------------------------------------------------------------- /Regexp/ENfa.agda: -------------------------------------------------------------------------------- 1 | open import Data.Nat using (ℕ; zero; suc) 2 | open import Data.Fin 3 | using (Fin) 4 | renaming (zero to fzero; suc to fsuc) 5 | open import Data.Fin.Subset as Subset 6 | using (Subset; ⁅_⁆; _∪_; _∩_; Nonempty) 7 | renaming (⊥ to ∅) 8 | open import Data.Bool using (false; true) 9 | open import Data.Vec using (toList) 10 | open import VecUtil 11 | 12 | module ENfa (Σ : Set) where 13 | open import String Σ using (String; _∷_; []) 14 | 15 | record eNfa (n : ℕ) : Set where 16 | field 17 | S : Fin n 18 | δ : Fin n → Σ → Subset n 19 | εδ : Fin n → Subset n 20 | F : Subset n 21 | 22 | eclose : ∀{n} → eNfa n → Subset n → Subset n 23 | eclose {n} enfa qs = ecloseQS enfa qs ∅ n 24 | where 25 | ecloseQ : {n : ℕ} → eNfa n → Subset n → Fin n → Subset n → ℕ → Subset n 26 | ecloseQS : {n : ℕ} → eNfa n → Subset n → Subset n → ℕ → Subset n 27 | 28 | ecloseQ enfa qs q visited m with visited ! q 29 | ... | true = ∅ 30 | ecloseQ enfa qs q visited zero | false = ∅ 31 | ecloseQ enfa qs q visited (suc m) | false = ⁅ q ⁆ ∪ (ecloseQS enfa (eNfa.εδ enfa q) (⁅ q ⁆ ∪ visited) m) 32 | 33 | ecloseQS enfa qs visited zero = ∅ 34 | ecloseQS enfa qs visited (suc m) = U (mapS qs (λ q → ecloseQ enfa qs q visited m) ∅) 35 | 36 | δ̂ : ∀{n} → eNfa n → (Subset n) → String → (Subset n) 37 | δ̂ {n} enfa qs [] = eclose enfa qs 38 | δ̂ {n} enfa qs (c ∷ s) = δ̂ enfa (onestep qs c) s 39 | where 40 | onestep : (Subset n) → Σ → (Subset n) 41 | onestep qs c = U (mapS (eclose enfa qs) (λ q → eNfa.δ enfa q c) ∅) 42 | 43 | 44 | infix 10 _↓_ 45 | _↓_ : ∀{n} → eNfa n → String → Set 46 | enfa ↓ s = Nonempty (eNfa.F enfa ∩ (δ̂ enfa ⁅ eNfa.S enfa ⁆ s)) 47 | -------------------------------------------------------------------------------- /report/latex/conclusion.tex: -------------------------------------------------------------------------------- 1 | \chapter{Conclusion} 2 | We formalized regular expressions and NFAs and proved the correctness of two different approaches for solving the matching problem. We showed that finite-state automata are closed under union, concatenation, and star operations and that the regular expressions language class is a subset of the NFAs language class. One further development could be to prove the inverse transformation. Our development ended with the formalization of the pumping lemma, which we used to prove the limits of these computation models. There is also a pumping lemma for the immediate superset of regular languages, context-free languages. The formalization of this class of languages could also be a future extension of the project.\\ 3 | 4 | An important conclusion that can be drawn from all this work is that sophisticated type-systems can be used as powerful tools that enforce writing correct code, to such an extent that compilers can be turned into theorem provers. Agda is a clear example of this.\\ 5 | On a proof assistant level, the constructive evidence approach used in Agda helps understand the flow of proofs flawlessly. In some cases, though, this also requires to provide constructions for many other seemingly obvious properties, which other theorem provers solve automatically. However, the Agda standard library is immense and it exposes plenty of well-known principles and properties. For instance, the file \texttt{Data.Nat.Properties} contains many properties about natural numbers and counts \texttt{2000+} lines of proofs.\\ 6 | When firstly approaching Agda, the learning curve can be very steep, but with a solid base in functional programming and Haskell, it becomes straightforward. -------------------------------------------------------------------------------- /Regexp/VecUtil.agda: -------------------------------------------------------------------------------- 1 | module VecUtil where 2 | open import Data.Nat using (ℕ; zero; suc) 3 | open import Data.Fin using (Fin; 0F) renaming (zero to fzero; suc to fsuc) 4 | open import Data.Vec using (Vec; lookup; tabulate; foldr; _∷_; []; _[_]=_; here; there) 5 | open import Data.Fin.Subset using (Subset; _∪_) renaming (⊥ to ∅) 6 | open import Data.Bool using (Bool; T; false; true; not) 7 | open import Relation.Nullary using (Dec; ¬_; yes; no) 8 | open import Data.Empty using (⊥; ⊥-elim) 9 | import Relation.Binary.PropositionalEquality as Eq 10 | open Eq using (_≡_; refl; subst; sym; trans; cong) 11 | open Eq.≡-Reasoning 12 | open import Data.Unit using (⊤; tt) 13 | 14 | !-syntax = lookup 15 | syntax !-syntax v i = v ! i 16 | 17 | ifPresentOrElse : ∀{m} {A : Set} → Fin m → Subset m → (B : Fin m → A) → A → A 18 | ifPresentOrElse i s f z with s ! i 19 | ... | false = z 20 | ... | true = f i 21 | 22 | mapS : {n : ℕ} {A : Set} → Subset n → (B : Fin n → A) → A → Vec A n 23 | mapS ss f z = tabulate λ i → ifPresentOrElse i ss f z 24 | 25 | U : ∀ {m}{n} → Vec (Subset n) m → Subset n 26 | U {n} = foldr _ _∪_ ∅ 27 | 28 | v[i]=v!i : ∀{n} {A : Set} 29 | → (V : Vec A n) 30 | → (i : Fin n) 31 | → V [ i ]= V ! i 32 | v[i]=v!i (x ∷ V) Data.Fin.0F = here 33 | v[i]=v!i (x ∷ V) (fsuc i) = there (v[i]=v!i V i) 34 | 35 | s!i≡s[i] : ∀{n}{A : Set}{s : Vec A n}{i}{v} 36 | → s [ i ]= v 37 | → s ! i ≡ v 38 | s!i≡s[i] here = refl 39 | s!i≡s[i] (there p) = s!i≡s[i] p 40 | 41 | 42 | _∈_ : ∀{n} → (p : Fin n) → (ss : Subset n) → Set 43 | p ∈ ss = T(lookup ss p) 44 | 45 | _∉_ : ∀{n} → (p : Fin n) → (ss : Subset n) → Set 46 | p ∉ ss = ¬ T(lookup ss p) 47 | 48 | _∈?_ : ∀{n} → (p : Fin n) → (ss : Subset n) → Dec (p ∈ ss) 49 | 0F ∈? (false ∷ z) = no (λ z → z) 50 | 0F ∈? (true ∷ z) = yes tt 51 | fsuc p ∈? (x ∷ z) = p ∈? z 52 | -------------------------------------------------------------------------------- /report/introduction_examples.agda: -------------------------------------------------------------------------------- 1 | module introduction_examples where 2 | open import Data.Nat 3 | open import Relation.Binary.PropositionalEquality 4 | 5 | data Even : ℕ → Set where 6 | zero : Even zero 7 | next : ∀{n} 8 | → Even n 9 | → Even (suc (suc n)) 10 | 11 | sum-is-even : ∀{n m} 12 | → Even n 13 | → Even m 14 | ------------ 15 | → Even (n + m) 16 | sum-is-even zero em = em 17 | sum-is-even (next en) em = next (sum-is-even en em) 18 | 19 | data _×_ (A B : Set) : Set where 20 | ⟨_,_⟩ : A → B → A × B 21 | 22 | data _⊎_ (A B : Set) : Set where 23 | inj₁ : A → A ⊎ B 24 | inj₂ : B → A ⊎ B 25 | 26 | data ⊥ : Set where 27 | 28 | data ⊤ : Set where 29 | tt : ⊤ 30 | 31 | ¬_ : Set → Set 32 | ¬ A = A → ⊥ 33 | 34 | ⊥-elim : ∀ {A : Set} 35 | → ⊥ 36 | → A 37 | ⊥-elim () 38 | 39 | contradiction : ∀{P Q} → P → ¬ P → Q 40 | contradiction p ¬p = ⊥-elim (¬p p) 41 | 42 | data Odd : ℕ → Set where 43 | one : Odd (suc zero) 44 | next : ∀{n} 45 | → Odd n 46 | → Odd (suc (suc n)) 47 | 48 | even⇒suc-odd : ∀{x} → Even x → Odd (suc x) 49 | even⇒suc-odd zero = one 50 | even⇒suc-odd (next a) = next (even⇒suc-odd a) 51 | 52 | odd⇒suc-even : ∀{x} → Odd x → Even (suc x) 53 | odd⇒suc-even one = next zero 54 | odd⇒suc-even (next a) = next (odd⇒suc-even a) 55 | 56 | ℕ-Even-or-Odd : ∀(x : ℕ) → Even x ⊎ Odd x 57 | ℕ-Even-or-Odd zero = inj₁ zero 58 | ℕ-Even-or-Odd (suc x) with ℕ-Even-or-Odd x 59 | ... | inj₁ ev = inj₂ (even⇒suc-odd ev) 60 | ... | inj₂ od = inj₁ (odd⇒suc-even od) 61 | 62 | even⇒¬odd : ∀{x : ℕ} → Even x → ¬ Odd x 63 | even⇒¬odd zero = λ () 64 | even⇒¬odd (next ev) (next od) = ⊥-elim((even⇒¬odd ev) od) 65 | 66 | ℕ-not-both-even-odd : ∀{x : ℕ} → ¬ (Even x × Odd x) 67 | ℕ-not-both-even-odd ⟨ ev , od ⟩ = ⊥-elim (even⇒¬odd ev od) 68 | 69 | data Dec (A : Set) : Set where 70 | yes : A → Dec A 71 | no : ¬ A → Dec A 72 | 73 | ¬even⇒suc-even : ∀{x : ℕ} → ¬ Even x → Even (suc x) 74 | ¬even⇒suc-even {zero} ne = ⊥-elim (ne zero) 75 | ¬even⇒suc-even {suc zero} ne = next zero 76 | ¬even⇒suc-even {suc (suc x)} ne = next (¬even⇒suc-even (λ z → ne (next z))) 77 | 78 | Even? : ∀(x : ℕ) → Dec (Even x) 79 | Even? x with ℕ-Even-or-Odd x 80 | ... | inj₁ xEven = yes xEven 81 | ... | inj₂ xOdd = no (λ xEven → ℕ-not-both-even-odd ⟨ xEven , xOdd ⟩) 82 | 83 | 84 | 85 | -- 86 | -------------------------------------------------------------------------------- /plfa/equality.agda: -------------------------------------------------------------------------------- 1 | module equality where 2 | 3 | data _≡_ {A : Set} (x : A) : A → Set where 4 | refl : x ≡ x 5 | 6 | infix 4 _≡_ 7 | 8 | sym : ∀ {A : Set} {x y : A} 9 | → x ≡ y 10 | ------- 11 | → y ≡ x 12 | sym refl = refl 13 | 14 | trans : ∀ {A : Set} {x y z : A} 15 | → x ≡ y 16 | → y ≡ z 17 | ------- 18 | → x ≡ z 19 | trans refl refl = refl 20 | 21 | cong : ∀ {A B : Set} (f : A → B) {x y : A} 22 | → x ≡ y 23 | ------- 24 | → f x ≡ f y 25 | cong f refl = refl 26 | 27 | cong₂ : ∀ {A B C : Set} (f : A → B → C) {u x : A} {v y : B} 28 | → u ≡ x 29 | → v ≡ y 30 | ------- 31 | → f u v ≡ f x y 32 | cong₂ f refl refl = refl 33 | 34 | cong-app : ∀ {A B : Set} {f g : A → B} 35 | → f ≡ g 36 | ------- 37 | → ∀ (x : A) → f x ≡ g x 38 | cong-app refl x = refl 39 | 40 | subst : ∀ {A : Set} {x y : A} (P : A → Set) 41 | → x ≡ y 42 | ------- 43 | → P x → P y 44 | subst P refl px = px 45 | 46 | -------------------------------- 47 | 48 | data ℕ : Set where 49 | zero : ℕ 50 | suc : ℕ → ℕ 51 | 52 | _+_ : ℕ → ℕ → ℕ 53 | zero + n = n 54 | (suc m) + n = suc (m + n) 55 | 56 | data _≤_ : ℕ → ℕ → Set where 57 | 58 | z≤n : ∀ {n : ℕ} 59 | -------- 60 | → zero ≤ n 61 | 62 | s≤s : ∀ {m n : ℕ} 63 | → m ≤ n 64 | ------------- 65 | → suc m ≤ suc n 66 | 67 | infix 4 _≤_ 68 | 69 | ≤-trans : ∀ {m n p : ℕ} 70 | → m ≤ n 71 | → n ≤ p 72 | ----- 73 | → m ≤ p 74 | ≤-trans z≤n _ = z≤n 75 | ≤-trans (s≤s m≤n) (s≤s n≤p) = s≤s (≤-trans m≤n n≤p) 76 | 77 | ≤-refl : ∀ {n : ℕ} 78 | ----- 79 | → n ≤ n 80 | ≤-refl {zero} = z≤n 81 | ≤-refl {suc n} = s≤s ≤-refl 82 | 83 | module ≤-Reasoning where 84 | 85 | infix 1 begin_ 86 | infixr 2 _≤⟨⟩_ _≤⟨_⟩_ 87 | infix 3 _∎ 88 | 89 | begin_ : ∀ {x y : ℕ} 90 | → x ≤ y 91 | ----- 92 | → x ≤ y 93 | begin x≤y = x≤y 94 | 95 | _≤⟨⟩_ : ∀ (x : ℕ) {y : ℕ} 96 | → x ≤ y 97 | ----- 98 | → x ≤ y 99 | x ≤⟨⟩ x≤y = x≤y 100 | 101 | _≤⟨_⟩_ : ∀ (x : ℕ) {y z : ℕ} 102 | → x ≤ y 103 | → y ≤ z 104 | ----- 105 | → x ≤ z 106 | x ≤⟨ x≤y ⟩ y≤z = ≤-trans x≤y y≤z 107 | 108 | _∎ : ∀ (x : ℕ) 109 | ----- 110 | → x ≤ x 111 | x ∎ = ≤-refl 112 | 113 | open ≤-Reasoning 114 | -------------------------------------------------------------------------------- /Regexp/DfaNfa.agda: -------------------------------------------------------------------------------- 1 | module DfaNfa (Σ : Set) where 2 | open import Data.Nat using (ℕ; suc; zero; _^_) 3 | open import Data.Fin using (Fin) renaming (zero to fzero; suc to fsuc) 4 | open import Data.Fin.Subset using (Subset; ⁅_⁆) renaming (⊥ to ∅) 5 | open import Data.Bool using (Bool; false; _∨_; _∧_; if_then_else_; T) 6 | open import Data.Vec using (lookup) 7 | open import Relation.Binary.PropositionalEquality using (_≡_; refl) 8 | open import Equivalence using (_⇔_; _≃_) 9 | open import String Σ using (String; []; _∷_) 10 | open import VecUtil using (U; mapS) 11 | open import SubsetFinIso renaming (subset-fin-iso to ISO) 12 | 13 | ∈ᵇ-syntax = lookup 14 | syntax ∈ᵇ-syntax v i = i ∈ᵇ v 15 | 16 | record DFA (n : ℕ) : Set₂ where 17 | field 18 | S : Fin n 19 | δ : Fin n → Σ → Fin n 20 | F : Fin n → Bool 21 | 22 | open DFA 23 | 24 | δd* : ∀{n} → DFA n → Fin n → String → Fin n 25 | δd* A q [] = q 26 | δd* A q (x ∷ xs) = δd* A (δ A q x) xs 27 | 28 | _↓ᵈ_ : ∀{n} → DFA n → String → Set 29 | A ↓ᵈ s = T (F A (δd* A (S A) s)) 30 | 31 | record Nfa (n : ℕ) : Set where 32 | field 33 | S : Fin n 34 | δ : Fin n → Σ → Subset n 35 | F : Fin n → Bool 36 | 37 | open Nfa 38 | 39 | δn* : ∀{n} → Nfa n → Subset n → String → Subset n 40 | δn* {n} A qs [] = qs 41 | δn* {n} A qs (c ∷ s) = δn* A (U (mapS qs (λ q → δ A q c) ∅)) s 42 | 43 | any : ∀{n} → (P : Fin n → Bool) → Bool 44 | any {zero} P = false 45 | any {suc _} P = P fzero ∨ any λ x → P (fsuc x) 46 | 47 | infix 10 _↓ⁿ_ 48 | _↓ⁿ_ : ∀{n} → Nfa n → String → Set 49 | A ↓ⁿ s = T(any λ q → F A q ∧ q ∈ᵇ (δn* A ⁅ S A ⁆ s)) 50 | 51 | NFAtoDFA : ∀{n} → Nfa n → DFA (2 ^ n) 52 | NFAtoDFA {n} N = record 53 | { S = _≃_.to ISO ⁅ S N ⁆ 54 | ; δ = λ qs c → _≃_.to ISO (U (mapS (_≃_.from ISO qs) (λ q → δ N q c) ∅)) 55 | ; F = λ qs → any λ q → F N q ∧ q ∈ᵇ _≃_.from ISO qs 56 | } 57 | 58 | prop : ∀{n} → (N : Nfa n) → ∀(qs) → (s : String) → 59 | δn* N qs s ≡ _≃_.from ISO (δd* (NFAtoDFA N) (_≃_.to ISO qs) s) 60 | prop N qs [] rewrite _≃_.from∘to ISO qs = refl 61 | prop N qs (x ∷ s) rewrite prop N (U (mapS qs (λ q → δ N q x) ∅)) s 62 | | _≃_.from∘to ISO qs = refl 63 | 64 | theorem1 : ∀{n : ℕ} → (N : Nfa n) → (s : String) → 65 | N ↓ⁿ s ⇔ NFAtoDFA N ↓ᵈ s 66 | theorem1 {n} N s rewrite prop N ⁅ S N ⁆ s = 67 | record { to = λ x → x ; from = λ x → x } 68 | -------------------------------------------------------------------------------- /report/latex/main.tex: -------------------------------------------------------------------------------- 1 | \documentclass[12pt]{book} 2 | 3 | \usepackage[utf8]{inputenc} 4 | \usepackage[english]{babel} 5 | \usepackage{amsmath,amssymb} 6 | \usepackage{amsthm} 7 | \usepackage{amsfonts} 8 | \usepackage[pdftex]{color, graphicx} 9 | \usepackage{makeidx} 10 | \usepackage{minted} 11 | \usepackage{newunicodechar} 12 | \usepackage{stmaryrd} 13 | \usepackage{geometry} 14 | \usepackage{times} 15 | 16 | \usepackage{tikz} 17 | \usetikzlibrary{automata,positioning} 18 | 19 | \usemintedstyle{bw} 20 | 21 | \newenvironment{agda} 22 | {% 23 | \VerbatimEnvironment 24 | \begin{minted}[fontsize=\small]{my-lexer.py:DAgdaLexer -x}% 25 | } 26 | {% 27 | \end{minted} 28 | } 29 | %%%%%%%%%%%%%%% BEGIN UNICODE SYMBOLS 30 | \newunicodechar{₀}{\ensuremath{_0}} 31 | \newunicodechar{₁}{\ensuremath{_1}} 32 | \newunicodechar{₂}{\ensuremath{_2}} 33 | \newunicodechar{₃}{\ensuremath{_3}} 34 | \newunicodechar{⁺}{\ensuremath{\textsuperscript{+}}} 35 | \newunicodechar{ˡ}{\ensuremath{^l}} 36 | \newunicodechar{ʳ}{\ensuremath{^r}} 37 | \newunicodechar{ᶠ}{\ensuremath{^f}} 38 | \newunicodechar{≡}{\ensuremath{\equiv}} 39 | \newunicodechar{≢}{\ensuremath{\not\equiv}} 40 | \newunicodechar{≟}{\mbox{\tiny\ensuremath{\stackrel{?}{=}}}} 41 | \newunicodechar{≈}{\ensuremath{\approx}} 42 | \newunicodechar{≃}{\ensuremath{\simeq}} 43 | \newunicodechar{≤}{\ensuremath{\le}} 44 | \newunicodechar{≥}{\ensuremath{\ge}} 45 | \newunicodechar{→}{\ensuremath{\rightarrow}} 46 | \newunicodechar{←}{\ensuremath{\leftarrow}} 47 | \newunicodechar{⇒}{\ensuremath{\Rightarrow}} 48 | \newunicodechar{⇔}{\ensuremath{\Leftrightarrow}} 49 | \newunicodechar{∀}{\ensuremath{\forall}} 50 | \newunicodechar{∃}{\ensuremath{\exists}} 51 | \newunicodechar{⟨}{\ensuremath{\langle}} 52 | \newunicodechar{⟩}{\ensuremath{\rangle}} 53 | \newunicodechar{⊤}{\ensuremath{\top}} 54 | \newunicodechar{⊎}{\ensuremath{\uplus}} 55 | \newunicodechar{×}{\ensuremath{\times}} 56 | \newunicodechar{∘}{\ensuremath{\circ}} 57 | \newunicodechar{ℕ}{\ensuremath{\mathbb{N}}} 58 | \newunicodechar{∈}{\ensuremath{\in}} 59 | \newunicodechar{¬}{\ensuremath{\neg}} 60 | \newunicodechar{⊥}{\ensuremath{\bot}} 61 | \newunicodechar{∩}{\ensuremath{\cap}} 62 | \newunicodechar{∪}{\ensuremath{\cup}} 63 | \newunicodechar{⊆}{\ensuremath{\subseteq}} 64 | \newunicodechar{ℓ}{\ensuremath{\ell}} 65 | \newunicodechar{∷}{\ensuremath{::}} 66 | \newunicodechar{Δ}{\ensuremath{\Delta}} 67 | \newunicodechar{Γ}{\ensuremath{\Gamma}} 68 | \newunicodechar{Σ}{\ensuremath{\Sigma}} 69 | \newunicodechar{Θ}{\ensuremath{\Theta}} 70 | \newunicodechar{Ω}{\ensuremath{\Omega}} 71 | \newunicodechar{ε}{\ensuremath{\epsilon}} 72 | \newunicodechar{λ}{\ensuremath{\lambda}} 73 | \newunicodechar{δ}{\ensuremath{\delta}} 74 | \newunicodechar{↓}{\ensuremath{\downarrow}} 75 | \newunicodechar{∅}{\ensuremath{\varnothing}} 76 | \newunicodechar{∁}{\ensuremath{\complement}} 77 | \newunicodechar{∉}{\ensuremath{\not \in}} 78 | \newunicodechar{⁆}{\ensuremath{\rrbracket}} 79 | \newunicodechar{⁅}{\ensuremath{\llbracket}} 80 | \newunicodechar{∎}{\ensuremath{\qed}} 81 | \newunicodechar{∧}{\ensuremath{\land}} 82 | \newunicodechar{∨}{\ensuremath{\lor}} 83 | \newunicodechar{⌊}{\ensuremath{\lfloor}} 84 | \newunicodechar{⌋}{\ensuremath{\rfloor}} 85 | \newunicodechar{⋃}{\ensuremath{\bigcup}} 86 | 87 | %%%%%%%%%%%%%%% END UNICODE SYMBOLS 88 | 89 | \newtheorem{theorem}{Theorem} 90 | \newtheorem{lemma}{Lemma} 91 | \DeclareMathOperator{\nullable}{Nullable} 92 | 93 | 94 | \begin{document} 95 | 96 | \include{first_page} 97 | 98 | \clearpage 99 | \newpage 100 | \thispagestyle{empty} 101 | 102 | \cleardoublepage 103 | 104 | \thispagestyle{empty} 105 | \newgeometry{centering} 106 | \section*{Abstract} 107 | We use the proof assistant Agda to prove properties on Regular Expressions and Finite State Automata. Initially, we prove the correctness of two ways to decide whether or not a string belongs to the language of a regular expression. The first one uses Brzozowski's Derivatives and the second one transforms regular expressions into Non-deterministic Finite State Automata. Finally, we show the pumping lemma for regular languages and conclude with the non-regularity of an example language. 108 | \thispagestyle{empty} 109 | \restoregeometry 110 | 111 | 112 | \thispagestyle{empty} 113 | \tableofcontents 114 | \thispagestyle{empty} 115 | 116 | \newpage 117 | \thispagestyle{empty} 118 | \clearpage\null\newpage 119 | \newpage 120 | \thispagestyle{empty} 121 | \clearpage\null\newpage 122 | 123 | \thispagestyle{empty} 124 | 125 | \addtocontents{toc}{\protect\thispagestyle{empty}} 126 | \setcounter{page}{1} 127 | \include{agda_introduction} 128 | 129 | \include{regular_expressions} 130 | 131 | \include{automata} 132 | 133 | \include{conclusion} 134 | 135 | \include{bibliography} 136 | 137 | \end{document} -------------------------------------------------------------------------------- /Regexp/SubsetFinIso.agda: -------------------------------------------------------------------------------- 1 | open import Relation.Binary.PropositionalEquality using (_≡_; _≗_; refl; sym; cong) 2 | open Relation.Binary.PropositionalEquality.≡-Reasoning 3 | open import Data.Fin using (Fin; zero; suc; inject+; raise) 4 | open import Data.Fin.Subset using (Subset; inside; outside) 5 | open import Data.Nat using (ℕ; suc; zero; _^_; _*_; _+_) 6 | open import Data.Nat.Properties using (+-identityʳ) 7 | open import Data.Fin.Properties using () 8 | open import Data.Vec using ([]; _∷_) 9 | open import Function using (_∘_; id) 10 | open import Data.Sum using (_⊎_; inj₁; inj₂; [_,_]; map) 11 | 12 | open import Equivalence 13 | 14 | [,]-∘-distr : ∀ {l} {A B C D : Set l} (f : A → B) 15 | {g : C → A} {h : D → A} → 16 | f ∘ [ g , h ] ≗ [ f ∘ g , f ∘ h ] 17 | [,]-∘-distr _ (inj₁ _) = refl 18 | [,]-∘-distr _ (inj₂ _) = refl 19 | 20 | [,]-map-commute : ∀ {l} {A B C D E : Set l} {f : A → B} {g : C → D} 21 | {f′ : B → E} {g′ : D → E} → 22 | [ f′ , g′ ] ∘ map f g ≗ [ f′ ∘ f , g′ ∘ g ] 23 | [,]-map-commute (inj₁ _) = refl 24 | [,]-map-commute (inj₂ _) = refl 25 | 26 | splitAt : ∀ m {n} → Fin (m + n) → Fin m ⊎ Fin n 27 | splitAt zero i = inj₂ i 28 | splitAt (suc m) zero = inj₁ zero 29 | splitAt (suc m) (suc i) = map suc id (splitAt m i) 30 | 31 | splitAt-inject+ : ∀ m n i → splitAt m (inject+ n i) ≡ inj₁ i 32 | splitAt-inject+ (suc m) n zero = refl 33 | splitAt-inject+ (suc m) n (suc i) rewrite splitAt-inject+ m n i = refl 34 | 35 | splitAt-raise : ∀ m n i → splitAt m (raise {n} m i) ≡ inj₂ i 36 | splitAt-raise zero n i = refl 37 | splitAt-raise (suc m) n i rewrite splitAt-raise m n i = refl 38 | 39 | inject+-raise-splitAt : ∀ m n i → [ inject+ n , raise {n} m ] (splitAt m i) ≡ i 40 | inject+-raise-splitAt zero n i = refl 41 | inject+-raise-splitAt (suc m) n zero = refl 42 | inject+-raise-splitAt (suc m) n (suc i) = begin 43 | [ inject+ n , raise {n} (suc m) ] (splitAt (suc m) (suc i)) ≡⟨ [,]-map-commute (splitAt m i) ⟩ 44 | [ suc ∘ (inject+ n) , suc ∘ (raise {n} m) ] (splitAt m i) ≡˘⟨ [,]-∘-distr suc (splitAt m i) ⟩ 45 | suc ([ inject+ n , raise {n} m ] (splitAt m i)) ≡⟨ cong suc (inject+-raise-splitAt m n i) ⟩ 46 | suc i ∎ 47 | 48 | sum-split : ∀ {m n} → 49 | Fin m ⊎ Fin n ≃ Fin (m + n) 50 | sum-split {m}{n} = record { to = to ; from = from ; from∘to = from∘to ; to∘from = to∘from } 51 | where 52 | to : Fin m ⊎ Fin n → Fin (m + n) 53 | to = [ inject+ n , raise {n} m ] 54 | 55 | from : Fin (m + n)→ Fin m ⊎ Fin n 56 | from = splitAt m 57 | 58 | from∘to : ∀ (x : Fin m ⊎ Fin n) → from (to x) ≡ x 59 | from∘to (inj₁ x) = splitAt-inject+ m n x 60 | from∘to (inj₂ y) = splitAt-raise m n y 61 | 62 | to∘from : ∀ (x : Fin (m + n)) → to (from x) ≡ x 63 | to∘from = inject+-raise-splitAt m n 64 | 65 | subset-suc : ∀ {n m} → 66 | Subset n ≃ Fin m 67 | → Subset (suc n) ≃ Fin m ⊎ Fin m 68 | subset-suc {n}{m} ev = record { to = to ; from = from ; from∘to = from∘to ; to∘from = to∘from } 69 | where 70 | to : Subset (suc n) → Fin m ⊎ Fin m 71 | to (outside ∷ r) = inj₁ (_≃_.to ev r) 72 | to (inside ∷ r) = inj₂ (_≃_.to ev r) 73 | 74 | from : Fin m ⊎ Fin m → Subset (suc n) 75 | from (inj₁ x) = outside ∷ (_≃_.from ev x) 76 | from (inj₂ y) = inside ∷ (_≃_.from ev y) 77 | 78 | from∘to : ∀ (x : Subset (suc n)) → from (to x) ≡ x 79 | from∘to (outside ∷ x) rewrite _≃_.from∘to ev x = refl 80 | from∘to (inside ∷ x) rewrite _≃_.from∘to ev x = refl 81 | 82 | to∘from : ∀ (x : Fin m ⊎ Fin m) → to (from x) ≡ x 83 | to∘from (inj₁ x) rewrite _≃_.to∘from ev x = refl 84 | to∘from (inj₂ y) rewrite _≃_.to∘from ev y = refl 85 | 86 | subset-fin-iso : ∀ {n} → Subset n ≃ Fin (2 ^ n) 87 | subset-fin-iso {zero} = record { to = λ _ → zero ; from = λ _ → [] ; from∘to = (λ { [] → refl }) ; to∘from = (λ { zero → refl }) } 88 | subset-fin-iso {suc n} with subset-fin-iso {n} 89 | ... | IH@record { to = to⁺ ; from = from⁺ ; from∘to = from∘to⁺ ; to∘from = to∘from⁺ } rewrite +-identityʳ (2 ^ n) = 90 | record { to = to ; from = from ; from∘to = from∘to ; to∘from = to∘from } 91 | where 92 | to : Subset (suc n) → Fin (2 ^ n + 2 ^ n) 93 | to s = _≃_.to sum-split (_≃_.to (subset-suc IH) s) 94 | 95 | from : Fin (2 ^ n + 2 ^ n) → Subset (suc n) 96 | from f = _≃_.from (subset-suc IH) (_≃_.from sum-split f) 97 | 98 | from∘to : ∀ (x : Subset (suc n)) → from (to x) ≡ x 99 | from∘to s rewrite _≃_.from∘to (sum-split {2 ^ n}{2 ^ n}) (_≃_.to (subset-suc IH) s) 100 | | _≃_.from∘to (subset-suc IH) s 101 | = refl 102 | 103 | to∘from : ∀ (x : Fin (2 ^ n + 2 ^ n)) → to (from x) ≡ x 104 | to∘from f rewrite _≃_.to∘from (subset-suc IH) (_≃_.from (sum-split{2 ^ n}{2 ^ n}) f) 105 | | _≃_.to∘from (sum-split {2 ^ n}{2 ^ n}) f 106 | = refl 107 | -------------------------------------------------------------------------------- /plfa/negation.agda: -------------------------------------------------------------------------------- 1 | module negation where 2 | import Relation.Binary.PropositionalEquality as Eq 3 | open Eq using (_≡_; refl) 4 | open Eq.≡-Reasoning using (begin_; _≡⟨⟩_; _≡⟨_⟩_; _∎) 5 | open import Data.Nat using (ℕ; zero; suc) 6 | open import Data.Empty using (⊥; ⊥-elim) 7 | open import Data.Sum using (_⊎_; inj₁; inj₂) 8 | open import isomorphism using (_≃_; extensionality) 9 | open import relations using (_<_) 10 | open import connectives using (⟨_,_⟩; _×_; proj₁; proj₂) 11 | 12 | ¬_ : Set → Set 13 | ¬ A = A → ⊥ 14 | 15 | ¬-elim : ∀ {A : Set} 16 | → ¬ A 17 | → A 18 | --- 19 | → ⊥ 20 | ¬-elim ¬x x = ¬x x 21 | 22 | infix 3 ¬_ 23 | 24 | ¬¬-intro : ∀ {A : Set} 25 | → A 26 | ----- 27 | → ¬ ¬ A 28 | ¬¬-intro x = λ{¬x → ¬x x} 29 | 30 | ¬¬¬-elim : ∀ {A : Set} 31 | → ¬ ¬ ¬ A 32 | ------- 33 | → ¬ A 34 | ¬¬¬-elim ¬¬¬x = λ x → ¬¬¬x (¬¬-intro x) 35 | 36 | contraposition : ∀ {A B : Set} 37 | → (A → B) 38 | ----------- 39 | → (¬ B → ¬ A) 40 | contraposition f ¬y x = ¬y (f x) 41 | 42 | _≢_ : ∀ {A : Set} → A → A → Set 43 | x ≢ y = ¬ (x ≡ y) 44 | 45 | _ : 1 ≢ 2 46 | _ = λ() 47 | 48 | peano : ∀ {m : ℕ} → zero ≢ suc m 49 | peano = λ() 50 | 51 | id : ⊥ → ⊥ 52 | id x = x 53 | 54 | id′ : ⊥ → ⊥ 55 | id′ () 56 | 57 | assimilation : ∀ {A : Set} (¬x ¬x′ : ¬ A) → ¬x ≡ ¬x′ 58 | assimilation ¬x ¬x′ = extensionality (λ x → ⊥-elim (¬x x)) 59 | 60 | -- Using negation, show that strict inequality is irreflexive, that is, n < n holds for no n. 61 | <-irreflexive : ∀ (n : ℕ) → ¬ (n < n) 62 | <-irreflexive (suc n) (_<_.s n 69 | -- Here “exactly one” means that not only one of the three must hold, 70 | -- but that when one holds the negation of the other two must also hold. 71 | trichotomy<₁ : ∀ (m n : ℕ) → m < n → ¬ (m ≡ n) 72 | trichotomy<₁ (suc m) (suc n) (_<_.sstable : ∀ {A : Set} → ¬ A → Stable A 161 | dne->stable a = λ x → ⊥-elim (x a) 162 | 163 | conj->stable : ∀ {A B : Set} → (Stable A) × (Stable B) → Stable ((Stable A) × (Stable B)) 164 | conj->stable z _ = z 165 | 166 | 167 | 168 | 169 | 170 | 171 | 172 | 173 | 174 | 175 | 176 | 177 | 178 | 179 | 180 | 181 | 182 | 183 | 184 | 185 | 186 | 187 | 188 | 189 | 190 | 191 | 192 | 193 | -- 194 | -------------------------------------------------------------------------------- /Regexp/Brzozowski.agda: -------------------------------------------------------------------------------- 1 | import Relation.Binary.PropositionalEquality as Eq 2 | open Eq using (_≡_; _≢_; refl; cong; sym; subst; trans) 3 | open import Relation.Nullary using (Dec; ¬_; yes; no) 4 | open import Relation.Nullary.Decidable using (map; ⌊_⌋) 5 | 6 | open import Data.Product using (_×_; Σ; ∃; Σ-syntax; ∃-syntax; proj₁; proj₂; _,_) 7 | open import Data.Empty using (⊥; ⊥-elim) 8 | open import Data.Sum using (_⊎_; inj₁; inj₂) 9 | open import Equivalence 10 | 11 | module Brzozowski (Σ : Set) (_≟_ : (a : Σ) → (b : Σ) → Dec (a ≡ b)) where 12 | open import Regexp Σ 13 | open import String Σ using (_++_; _∷_; ++-assoc; []; String; ++-idʳ; ++-idˡ; foldl) 14 | 15 | open import Data.Bool using (if_then_else_) 16 | 17 | data Nullable : RegExp → Set where 18 | null⟨ε⟩ : Nullable ⟨ε⟩ 19 | null+l : ∀{F G} → Nullable F → Nullable (F + G) 20 | null+r : ∀{F G} → Nullable G → Nullable (F + G) 21 | null· : ∀{F G} → Nullable F → Nullable G → Nullable (F · G) 22 | null* : ∀{F} → Nullable (F *) 23 | 24 | Nullable? : (E : RegExp) → Dec (Nullable E) 25 | Nullable? ⟨⟩ = no (λ ()) 26 | Nullable? ⟨ε⟩ = yes null⟨ε⟩ 27 | Nullable? (Atom c) = no (λ ()) 28 | Nullable? (r + s) with Nullable? r | Nullable? s 29 | ... | yes p | _ = yes (null+l p) 30 | ... | _ | yes q = yes (null+r q) 31 | ... | no ¬p | no ¬q = no λ{ (null+l p) → ⊥-elim (¬p p) 32 | ; (null+r q) → ⊥-elim (¬q q) } 33 | Nullable? (r · s) with Nullable? r | Nullable? s 34 | ... | yes p | yes q = yes (null· p q) 35 | ... | _ | no ¬q = no λ{ (null· _ q) → ⊥-elim (¬q q) } 36 | ... | no ¬p | _ = no λ{ (null· p _) → ⊥-elim (¬p p) } 37 | Nullable? (r *) = yes null* 38 | 39 | _[_] : RegExp → Σ → RegExp 40 | ⟨⟩ [ a ] = ⟨⟩ 41 | ⟨ε⟩ [ a ] = ⟨⟩ 42 | (Atom b)[ a ] = if ⌊ b ≟ a ⌋ 43 | then ⟨ε⟩ 44 | else ⟨⟩ 45 | (F + G)[ a ] = F [ a ] + G [ a ] 46 | (F · G)[ a ] = if ⌊ Nullable? F ⌋ 47 | then F [ a ] · G + G [ a ] 48 | else F [ a ] · G 49 | (F *)[ a ] = F [ a ] · F * 50 | 51 | split-seq : ∀{s E F} 52 | → s ∈ (E · F) 53 | → ∃[ u ] ∃[ v ] ((s ≡ u ++ v) × (u ∈ E) × (v ∈ F)) 54 | split-seq (in-· p q) = _ , _ , refl , p , q 55 | 56 | ε-seq : ∀{E F} → ε ∈ (E · F) → ε ∈ E × ε ∈ F 57 | ε-seq p with split-seq p 58 | ... | [] , [] , refl , p1 , p2 = p1 , p2 59 | 60 | split-* : ∀{E s} 61 | → s ∈ (E *) 62 | → s ≢ ε 63 | → ∃[ u ] ∃[ v ] (u ≢ ε × s ≡ u ++ v × u ∈ E × v ∈ (E *)) 64 | split-* in-*1 q = ⊥-elim (q refl) 65 | split-* (in-*2 {[]} p q) neps = split-* q neps 66 | split-* (in-*2 {x ∷ s} {t} p q) _ = x ∷ s , t , (λ ()) , refl , p , q 67 | 68 | theorem1 : ∀{E : RegExp} 69 | → ε ∈(E) ⇔ Nullable E 70 | theorem1 = record { to = to ; from = from } 71 | where 72 | to : ∀{E} → ε ∈ E → Nullable E 73 | to {⟨ε⟩} _ = null⟨ε⟩ 74 | to {E + F} (in+l x) = null+l (to x) 75 | to {E + F} (in+r x) = null+r (to x) 76 | to {E · F} x with ε-seq x 77 | ... | ε∈E , ε∈F = null· (to ε∈E) (to ε∈F) 78 | to {E *} _ = null* 79 | 80 | from : ∀{E} → Nullable E → ε ∈ E 81 | from null⟨ε⟩ = in-ε 82 | from (null+l x) = in+l (from x) 83 | from (null+r x) = in+r (from x) 84 | from (null· x y) = in-· (from x) (from y) 85 | from null* = in-*1 86 | 87 | theorem2 : ∀{a : Σ} {v : String} {E : RegExp} 88 | → v ∈(E [ a ]) ⇔ (a ∷ v) ∈(E) 89 | theorem2 = record { to = to ; from = from } 90 | where 91 | to : ∀{a v E} → v ∈(E [ a ]) → (a ∷ v) ∈( E ) 92 | to {a} {v} {Atom c} x with c ≟ a 93 | to {_} {[]} {Atom c} x | yes refl = in-c c 94 | to {E = F + G} (in+l x) = in+l (to x) 95 | to {E = F + G} (in+r x) = in+r (to x) 96 | to {E = F · G} x with Nullable? F 97 | to {E = F · G} (in+l (in-· x y)) | yes p 98 | = in-· (to x) y 99 | to {E = F · G} (in+r x) | yes p 100 | = in-· (_⇔_.from theorem1 p) (to x) 101 | to {E = F · G} (in-· x y) | no ¬p 102 | = in-· (to x) y 103 | to {E = F *} (in-· x y) = in-*2 (to x) y 104 | 105 | from : ∀ {a}{v}{E} → (a ∷ v) ∈ E → v ∈ E [ a ] 106 | from {_} {_} {Atom c} (in-c .c) with c ≟ c 107 | ... | yes p = in-ε 108 | ... | no ¬p = ⊥-elim (¬p refl) 109 | from {E = F + G} (in+l x) = in+l (from x) 110 | from {E = F + G} (in+r x) = in+r (from x) 111 | from {E = F · G} x with Nullable? F | split-seq x 112 | ... | yes p | [] , av , refl , _ , av∈G 113 | = in+r (from av∈G) 114 | ... | yes p | a ∷ u , t , refl , au∈F , t∈G 115 | = in+l (in-· (from au∈F) t∈G) 116 | ... | no ¬p | [] , _ , refl , ε∈F , _ 117 | = ⊥-elim (¬p (_⇔_.to theorem1 ε∈F)) 118 | ... | no ¬p | a ∷ u , t , refl , au∈F , t∈G 119 | = in-· (from au∈F) t∈G 120 | from {E = F *} x with split-* x (λ ()) 121 | ... | [] , _ , ¬p , _ , _ , _ = ⊥-elim (¬p refl) 122 | ... | a ∷ t , v , _ , refl , at∈E , v∈E* 123 | = in-· (from at∈E) v∈E* 124 | 125 | Iff⇒ = _⇔_.to 126 | Iff⇐ = _⇔_.from 127 | 128 | theorem3 : ∀ {v F} → v ∈(F) ⇔ Nullable (foldl _[_] F v) 129 | theorem3 = record { to = to ; from = from } 130 | where 131 | to : ∀ {v F} → v ∈(F) → Nullable (foldl _[_] F v) 132 | to {[]} x = Iff⇒ theorem1 x 133 | to {v ∷ vs} x = to (Iff⇐ theorem2 x) 134 | 135 | from : ∀ {v F} → Nullable (foldl _[_] F v) → v ∈ F 136 | from {[]} x = Iff⇐ theorem1 x 137 | from {v ∷ vs} x = Iff⇒ theorem2 (from x) 138 | 139 | _∈?_ : (v : String) → (F : RegExp) → Dec (v ∈ F) 140 | v ∈? F with Nullable? (foldl _[_] F v) 141 | ... | yes p = yes (Iff⇐ theorem3 p) 142 | ... | no ¬p = no (λ z → ¬p (Iff⇒ theorem3 z)) 143 | -------------------------------------------------------------------------------- /plfa/decidable.agda: -------------------------------------------------------------------------------- 1 | module decidable where 2 | import Relation.Binary.PropositionalEquality as Eq 3 | open Eq using (_≡_; refl) 4 | open Eq.≡-Reasoning 5 | open import Data.Nat using (ℕ; zero; suc) 6 | open import Data.Product using (_×_) renaming (_,_ to ⟨_,_⟩) 7 | open import Data.Sum using (_⊎_; inj₁; inj₂) 8 | open import Relation.Nullary using (¬_) 9 | open import Relation.Nullary.Negation using () 10 | renaming (contradiction to ¬¬-intro) 11 | open import Data.Unit using (⊤; tt) 12 | open import Data.Empty using (⊥; ⊥-elim) 13 | open import relations using (_<_; _≤_; z≤n; s≤s; z m + zero ≡ m 30 | +-identity zero = 31 | begin 32 | zero + zero 33 | ≡⟨⟩ 34 | zero 35 | ∎ 36 | +-identity (suc n) = 37 | begin 38 | suc n + zero 39 | ≡⟨⟩ 40 | suc (n + zero) 41 | ≡⟨ cong suc (+-identity n) ⟩ 42 | suc n 43 | ∎ 44 | 45 | +-suc : ∀ (m n : ℕ) -> m + suc n ≡ suc (m + n) 46 | +-suc zero n = 47 | begin 48 | zero + suc n 49 | ≡⟨⟩ 50 | suc n 51 | ∎ 52 | +-suc (suc m) n = 53 | begin 54 | suc m + suc n 55 | ≡⟨⟩ 56 | suc (m + suc n) 57 | ≡⟨ cong suc (+-suc m n) ⟩ 58 | suc (suc (m + n)) 59 | ∎ 60 | 61 | +-comm : ∀ (m n : ℕ) → m + n ≡ n + m 62 | +-comm m zero = 63 | begin 64 | m + zero 65 | ≡⟨ +-identity m ⟩ 66 | m 67 | ≡⟨⟩ 68 | zero + m 69 | ∎ 70 | +-comm m (suc n) = 71 | begin 72 | m + suc n 73 | ≡⟨ +-suc m n ⟩ 74 | suc (m + n) 75 | ≡⟨ cong suc (+-comm m n) ⟩ 76 | suc (n + m) 77 | ≡⟨⟩ 78 | suc n + m 79 | ∎ 80 | 81 | +-rearrange : ∀ (m n p q : ℕ) → (m + n) + (p + q) ≡ m + (n + p) + q 82 | +-rearrange m n p q = 83 | begin 84 | (m + n) + (p + q) 85 | ≡⟨ +-assoc m n (p + q) ⟩ 86 | m + (n + (p + q)) 87 | ≡⟨ cong (m +_) (sym (+-assoc n p q)) ⟩ 88 | m + ((n + p) + q) 89 | ≡⟨ sym (+-assoc m (n + p) q) ⟩ 90 | (m + (n + p)) + q 91 | ∎ 92 | 93 | +-assocRewrite : ∀ (m n p : ℕ) → m + (n + p) ≡ (m + n) + p 94 | +-assocRewrite zero n p = refl 95 | +-assocRewrite (suc m) n p rewrite +-assocRewrite m n p = refl 96 | 97 | +-swap : ∀ (m n p : ℕ) → m + (n + p) ≡ n + (m + p) 98 | +-swap m n p = 99 | begin 100 | m + (n + p) 101 | ≡⟨ sym (+-assoc m n p) ⟩ 102 | (m + n) + p 103 | ≡⟨ cong (_+ p) (+-comm m n) ⟩ 104 | (n + m) + p 105 | ≡⟨ +-assoc n m p ⟩ 106 | n + (m + p) 107 | ∎ 108 | 109 | *-distrib-+ : ∀ (m n p : ℕ) → (m + n) * p ≡ m * p + n * p 110 | *-distrib-+ zero n p = refl 111 | *-distrib-+ (suc m) n p rewrite *-distrib-+ m n p = sym (+-assoc p (m * p) (n * p)) 112 | 113 | *-assoc : ∀ (m n p : ℕ) → (m * n) * p ≡ m * (n * p) 114 | *-assoc zero n p = refl 115 | *-assoc (suc m) n p rewrite *-distrib-+ n (m * n) p | *-assoc m n p = refl 116 | 117 | 118 | *-id : ∀ (n : ℕ) → n * (suc 0) ≡ n 119 | *-id zero = refl 120 | *-id (suc n) rewrite *-id n = refl 121 | 122 | *-nullo : ∀ (n : ℕ) → n * zero ≡ zero 123 | *-nullo zero = refl 124 | *-nullo (suc n) rewrite *-nullo n = refl 125 | 126 | *-uno : ∀ (n m : ℕ) → n + n * m ≡ n * suc m 127 | *-uno zero m rewrite *-nullo m | *-id m | +-identity m = refl 128 | *-uno (suc n) m = 129 | begin 130 | (suc n) + (suc n * m) 131 | ≡⟨⟩ 132 | suc (n + suc n * m) 133 | ≡⟨⟩ 134 | suc (n + suc n * m) 135 | ≡⟨⟩ 136 | suc (n + (m + n * m)) 137 | ≡⟨ cong suc (sym (+-assoc n m (n * m))) ⟩ 138 | suc ((n + m) + n * m) 139 | ≡⟨ cong suc (cong (_+ n * m) (+-comm n m)) ⟩ 140 | suc ((m + n) + n * m) 141 | ≡⟨ cong suc (+-assoc m n (n * m)) ⟩ 142 | suc (m + (n + n * m)) 143 | ≡⟨ cong suc (cong (m +_) (*-uno n m)) ⟩ 144 | suc (m + n * suc m) 145 | ≡⟨⟩ 146 | suc m + n * suc m 147 | ≡⟨⟩ 148 | suc n * suc m 149 | ∎ 150 | 151 | *-comm : ∀ (m n : ℕ) → m * n ≡ n * m 152 | *-comm zero n rewrite *-nullo n = refl 153 | *-comm (suc m) n rewrite *-comm m n | *-uno n m = refl 154 | 155 | 0∸n≡0 : ∀ (n : ℕ) → zero ∸ n ≡ zero 156 | 0∸n≡0 zero = refl 157 | 0∸n≡0 (suc n) = refl 158 | 159 | ∸-+-assoc : ∀ (m n p : ℕ) → m ∸ n ∸ p ≡ m ∸ (n + p) 160 | ∸-+-assoc zero n p rewrite 0∸n≡0 (n + p) | 0∸n≡0 n | 0∸n≡0 p = refl 161 | ∸-+-assoc (suc m) zero p = refl 162 | ∸-+-assoc (suc m) (suc n) p rewrite ∸-+-assoc m n p = refl 163 | 164 | data Bin : Set where 165 | ⟨⟩ : Bin 166 | _O : Bin → Bin 167 | _I : Bin → Bin 168 | 169 | inc : Bin → Bin 170 | inc ⟨⟩ = ⟨⟩ I 171 | inc (prec O) = prec I 172 | inc (prec I) = (inc prec) O 173 | 174 | natToBin : ℕ → Bin 175 | natToBin zero = ⟨⟩ O 176 | natToBin (suc n) = inc (natToBin n) 177 | 178 | binToNat : Bin → ℕ 179 | binToNat ⟨⟩ = 0 180 | binToNat (b O) = 2 * binToNat b 181 | binToNat (b I) = suc (2 * binToNat b) 182 | 183 | lemma1 : ∀ (b : Bin) → suc (binToNat (b I)) ≡ 2 + binToNat b * 2 184 | lemma1 ⟨⟩ = refl 185 | lemma1 (b O) = 186 | begin 187 | suc (binToNat ((b O) I)) 188 | ≡⟨⟩ 189 | suc (suc (2 * (binToNat (b O)))) 190 | ≡⟨ cong suc (cong suc (*-comm 2 (binToNat (b O)))) ⟩ 191 | suc (suc (binToNat (b O) * 2)) 192 | ≡⟨⟩ 193 | 2 + binToNat (b O) * 2 194 | ∎ 195 | lemma1 (b I) = 196 | begin 197 | suc (binToNat ((b I) I)) 198 | ≡⟨⟩ 199 | suc (suc (2 * binToNat (b I))) 200 | ≡⟨ cong suc (cong suc (*-comm 2 (binToNat (b I)))) ⟩ 201 | 2 + binToNat (b I) * 2 202 | ∎ 203 | 204 | binLaw1 : ∀ (b : Bin) → binToNat (inc b) ≡ suc (binToNat b) 205 | binLaw1 ⟨⟩ = refl 206 | binLaw1 (b O) = refl 207 | binLaw1 (b I) = 208 | begin 209 | binToNat (inc (b I)) 210 | ≡⟨⟩ 211 | binToNat (inc b O) 212 | ≡⟨⟩ 213 | 2 * binToNat (inc b) 214 | ≡⟨ cong (2 *_) (binLaw1 (b)) ⟩ 215 | 2 * (suc (binToNat b)) 216 | ≡⟨ *-comm 2 (suc (binToNat b))⟩ 217 | (suc (binToNat b)) * 2 218 | ≡⟨ *-distrib-+ 1 (binToNat b) 2 ⟩ 219 | 2 + ((binToNat b) * 2) 220 | ≡⟨ sym (lemma1 b) ⟩ 221 | suc (binToNat (b I)) 222 | ∎ 223 | 224 | binLaw2 : ∀ (n : ℕ) → binToNat (natToBin n) ≡ n 225 | binLaw2 zero = refl 226 | binLaw2 (suc n) = 227 | begin 228 | binToNat (natToBin (suc n)) 229 | ≡⟨⟩ 230 | binToNat (inc (natToBin n)) 231 | ≡⟨ binLaw1 (natToBin n) ⟩ 232 | suc (binToNat (natToBin n)) 233 | ≡⟨ cong suc (binLaw2 n) ⟩ 234 | suc n 235 | ∎ 236 | 237 | -- Claim: natToBin (binToNat b) ≡ b 238 | -- does not hold: natToBin (binToNat ⟨⟩) = ⟨⟩ O and so ⟨⟩ O /= ⟨⟩ 239 | -------------------------------------------------------------------------------- /Regexp/RegexpNfa.agda: -------------------------------------------------------------------------------- 1 | open import Equivalence 2 | open import Data.Nat as ℕ using (ℕ; zero; suc; _<′_; ≤′-refl; ≤′-step) 3 | open import Data.Nat.Properties 4 | open import Data.Fin 5 | using (Fin; toℕ; inject+; raise; 0F; 1F; 2F; 3F; 4F; 5F; 6F) 6 | renaming (_<_ to _<ᶠ_; zero to fzero; suc to fsuc) 7 | open import Data.Fin.Subset as Subset 8 | using (Subset; ⁅_⁆) 9 | renaming (⊥ to ∅; ⊤ to FullSet) 10 | open import Relation.Nullary using (Dec; yes; no; ¬_) 11 | open import Data.Bool using (Bool; false; true; not; T; _∨_) 12 | open import Data.Empty as Empty using (⊥; ⊥-elim) 13 | open import Relation.Nullary.Negation using (contradiction) 14 | open import Data.Unit using (⊤; tt) 15 | open import Data.Product using (_×_; ∃; ∃₂; Σ-syntax; ∃-syntax; _,_; proj₁; proj₂) 16 | open import Data.Sum using (_⊎_; inj₁; inj₂) 17 | import Relation.Binary.PropositionalEquality as Eq 18 | open Eq using (_≡_; refl; _≢_; subst; sym; trans; cong) 19 | open import VecUtil hiding (_∈?_; _∈_) 20 | open import Induction.Nat 21 | open import Induction.WellFounded 22 | 23 | module RegexpNfa (Σ : Set) (_≟_ : (a : Σ) → (b : Σ) → Dec (a ≡ b)) where 24 | open import String Σ using (String; _∷_; []; length) renaming (_++_ to _++ˢ_) 25 | open import Nfa Σ hiding (Acc; ε) 26 | open import Regexp Σ 27 | 28 | nfa-∅ : Nfa 1 29 | nfa-∅ = record { S = 0F ; δ = λ _ _ → ⁅ 0F ⁆ ; F = ∅ } 30 | 31 | nfa-ε : Nfa 2 32 | nfa-ε = record { S = 0F ; δ = λ _ _ → ⁅ 1F ⁆ ; F = ⁅ 0F ⁆ } 33 | 34 | nfa-c : (c : Σ) → Nfa 3 35 | nfa-c c = record { S = 0F ; δ = δ ; F = ⁅ 1F ⁆ } 36 | where 37 | δ : Fin 3 → Σ → Subset 3 38 | δ 0F k with k ≟ c 39 | ... | yes p = ⁅ 1F ⁆ 40 | ... | no ¬p = ⁅ 2F ⁆ 41 | δ _ _ = ⁅ 2F ⁆ 42 | 43 | extractOrL : ∀{a} → T(a ∨ false) → T a 44 | extractOrL {false} z = z 45 | extractOrL {true} _ = tt 46 | 47 | nfa-∅-is-empty : (s : String) → ¬ (nfa-∅ ↓ s) 48 | nfa-∅-is-empty (x ∷ s) r = ⊥-elim (nfa-∅-is-empty s (extractOrL r)) 49 | 50 | 1F-is-error : (s : String) → ¬ (T (accepts nfa-ε 1F s)) 51 | 1F-is-error [] z = z 52 | 1F-is-error (x ∷ s) z = ⊥-elim (1F-is-error s (extractOrL z)) 53 | 54 | nfaε-correct : (s : String) → ¬ (s ≡ ε) → ¬ (nfa-ε ↓ s) 55 | nfaε-correct [] a b = a refl 56 | nfaε-correct (x ∷ s) a b = 1F-is-error s (extractOrL b) 57 | 58 | 2F-is-error : ∀{c s} → ¬ (T (accepts (nfa-c c) 2F s )) 59 | 2F-is-error {c} {x ∷ s} d = 60 | contradiction (extractOrL d) (2F-is-error {c} {s}) 61 | 62 | nfac-correct : ∀{c}{s} → nfa-c c ↓ s → s ≡ (c ∷ []) 63 | nfac-correct {c} {x ∷ []} d with x ≟ c 64 | ... | yes p = cong (_∷ []) p 65 | nfac-correct {c} {x ∷ y ∷ s} d with x ≟ c 66 | ... | yes p = contradiction d (2F-is-error {c} {x ∷ y ∷ s}) 67 | ... | no ¬p = contradiction d (2F-is-error {c} {x ∷ y ∷ s}) 68 | 69 | 70 | toNFA : (R : RegExp) 71 | → ∃₂ λ (n : ℕ) (nfa : Nfa n) 72 | → ∀ (s : String) 73 | → s ∈ R ⇔ nfa ↓ s 74 | toNFA ⟨⟩ = 1 , nfa-∅ , λ s → record 75 | { to = λ () 76 | ; from = λ nfa↓s → ⊥-elim (nfa-∅-is-empty s nfa↓s) 77 | } 78 | 79 | toNFA ⟨ε⟩ = 2 , nfa-ε , iff 80 | where 81 | iff : (s : String) 82 | → s ∈ ⟨ε⟩ ⇔ nfa-ε ↓ s 83 | iff [] = record { to = λ _ → tt ; from = λ _ → in-ε } 84 | iff (x ∷ s) = record 85 | { to = λ () 86 | ; from = λ nfa↓xs → ⊥-elim (nfaε-correct (x ∷ s) (λ ()) nfa↓xs ) 87 | } 88 | 89 | toNFA (Atom c) = 3 , nfa-c c , λ s → to IFF from 90 | where 91 | to : ∀{s} 92 | → s ∈ Atom c 93 | → nfa-c c ↓ s 94 | to (in-c c) with c ≟ c 95 | ... | yes p = tt 96 | ... | no ¬p = ¬p refl 97 | 98 | from : ∀{s} 99 | → nfa-c c ↓ s 100 | → s ∈ Atom c 101 | from {s} nfa↓s rewrite nfac-correct {c} {s} nfa↓s = in-c c 102 | 103 | toNFA (R + F) with toNFA R | toNFA F 104 | ... | n , A , w∈R⇔A↓w | m , B , w∈F⇔B↓w = 105 | suc n ℕ.+ m , union A B , λ s → to s IFF (from s) 106 | where 107 | to : (s : String) 108 | → s ∈ (R + F) 109 | → union A B ↓ s 110 | to s (in+l s∈R) 111 | = _⇔_.to (union-correct s) (inj₁ (_⇔_.to (w∈R⇔A↓w s) s∈R)) 112 | to s (in+r s∈F) 113 | = _⇔_.to (union-correct s) (inj₂ (_⇔_.to (w∈F⇔B↓w s) s∈F)) 114 | 115 | from : (s : String) 116 | → union A B ↓ s 117 | → s ∈ (R + F) 118 | from s A∪B↓s with _⇔_.from (union-correct s) A∪B↓s 119 | ...| inj₁ A↓s = in+l (_⇔_.from (w∈R⇔A↓w s) A↓s) 120 | ...| inj₂ B↓s = in+r (_⇔_.from (w∈F⇔B↓w s) B↓s) 121 | 122 | toNFA (R · F) with toNFA R | toNFA F 123 | ... | n , A , w∈R⇔A↓w | m , B , w∈F⇔B↓w = 124 | suc n ℕ.+ m , concat A B , λ s → to s IFF (from s) 125 | where 126 | to : (s : String) 127 | → s ∈ (R · F) 128 | → concat A B ↓ s 129 | to _ (in-· {u} {v} u∈R v∈F) = 130 | _⇔_.from (concat-correct (u ++ˢ v)) 131 | (u 132 | , v 133 | , refl 134 | , _⇔_.to (w∈R⇔A↓w u) u∈R 135 | , _⇔_.to (w∈F⇔B↓w v) v∈F) 136 | 137 | from : (s : String) 138 | → concat A B ↓ s 139 | → s ∈ (R · F) 140 | from s AB↓s with _⇔_.to (concat-correct s) AB↓s 141 | ... | u , v , s≡uv , A↓u , B↓v rewrite s≡uv = 142 | in-· (_⇔_.from (w∈R⇔A↓w u) A↓u) (_⇔_.from (w∈F⇔B↓w v) B↓v) 143 | 144 | toNFA (R *) with toNFA R 145 | ... | n , A , s∈R⇔A↓s = 146 | suc n , star A , λ s → (to s) IFF (from s) 147 | where 148 | to : (s : String) 149 | → s ∈ (R *) 150 | → star A ↓ s 151 | to _ in-*1 = tt 152 | to _ (in-*2 {u} {v} u∈R v∈R*) = 153 | star-correct1 {_} {u} {v} (_⇔_.to (s∈R⇔A↓s u) u∈R , (to v v∈R*)) 154 | 155 | lenv?@^|_~:\\]).*?$', Comment.Single), 38 | (r'\{-', Comment.Multiline, 'comment'), 39 | # Lexemes: 40 | # Identifiers 41 | (r'\bimport\b', Keyword.Reserved, 'import'), 42 | (r'\bmodule\b', Keyword.Reserved, 'module'), 43 | (r'\berror\b', Name.Exception), 44 | (r'\b(%s)(?!\')\b' % '|'.join(reserved), Keyword.Reserved), 45 | (r"'[^\\]'", String.Char), # this has to come before the TH quote 46 | (r'^[_' + uni.Ll + r'][\w\']*', Name.Function), 47 | (r"'?[_" + uni.Ll + r"][\w']*", Name), 48 | (r"('')?[" + uni.Lu + r"][\w\']*", Keyword.Type), 49 | (r"(')[" + uni.Lu + r"][\w\']*", Keyword.Type), 50 | (r"(')\[[^\]]*\]", Keyword.Type), # tuples and lists get special treatment in GHC 51 | (r"(')\([^)]*\)", Keyword.Type), # .. 52 | (r"(')[:!#$%&*+.\\/<=>?@^|~-]+", Keyword.Type), # promoted type operators 53 | # Operators 54 | (r'\\(?![:!#$%&*+.\\/<=>?@^|~-]+)', Name.Function), # lambda operator 55 | (r'(<-|::|->|=>|=)(?![:!#$%&*+.\\/<=>?@^|~-]+)', Operator.Word), # specials 56 | (r':[:!#$%&*+.\\/<=>?@^|~-]*', Keyword.Type), # Constructor operators 57 | (r'[:!#$%&*+.\\/<=>?@^|~-]+', Operator), # Other operators 58 | # Numbers 59 | (r'0[xX]_*[\da-fA-F](_*[\da-fA-F])*_*[pP][+-]?\d(_*\d)*', Number.Float), 60 | (r'0[xX]_*[\da-fA-F](_*[\da-fA-F])*\.[\da-fA-F](_*[\da-fA-F])*' 61 | r'(_*[pP][+-]?\d(_*\d)*)?', Number.Float), 62 | (r'\d(_*\d)*_*[eE][+-]?\d(_*\d)*', Number.Float), 63 | (r'\d(_*\d)*\.\d(_*\d)*(_*[eE][+-]?\d(_*\d)*)?', Number.Float), 64 | (r'0[bB]_*[01](_*[01])*', Number.Bin), 65 | (r'0[oO]_*[0-7](_*[0-7])*', Number.Oct), 66 | (r'0[xX]_*[\da-fA-F](_*[\da-fA-F])*', Number.Hex), 67 | (r'\d(_*\d)*', Number.Integer), 68 | # Character/String Literals 69 | (r"'", String.Char, 'character'), 70 | (r'"', String, 'string'), 71 | # Special 72 | (r'\[\]', Keyword.Type), 73 | (r'\(\)', Name.Builtin), 74 | (r'[][(),;`{}]', Punctuation), 75 | ], 76 | 'import': [ 77 | # Import statements 78 | (r'\s+', Text), 79 | (r'"', String, 'string'), 80 | # after "funclist" state 81 | (r'\)', Punctuation, '#pop'), 82 | (r'qualified\b', Keyword), 83 | # import X as Y 84 | (r'([' + uni.Lu + r'][\w.]*)(\s+)(as)(\s+)([' + uni.Lu + r'][\w.]*)', 85 | bygroups(Name.Namespace, Text, Keyword, Text, Name), '#pop'), 86 | # import X hiding (functions) 87 | (r'([' + uni.Lu + r'][\w.]*)(\s+)(hiding)(\s+)(\()', 88 | bygroups(Name.Namespace, Text, Keyword, Text, Punctuation), 'funclist'), 89 | # import X (functions) 90 | (r'([' + uni.Lu + r'][\w.]*)(\s+)(\()', 91 | bygroups(Name.Namespace, Text, Punctuation), 'funclist'), 92 | # import X 93 | (r'[\w.]+', Name.Namespace, '#pop'), 94 | ], 95 | 'module': [ 96 | (r'\s+', Text), 97 | (r'([' + uni.Lu + r'][\w.]*)(\s+)(\()', 98 | bygroups(Name.Namespace, Text, Punctuation), 'funclist'), 99 | (r'[' + uni.Lu + r'][\w.]*', Name.Namespace, '#pop'), 100 | ], 101 | 'funclist': [ 102 | (r'\s+', Text), 103 | (r'[' + uni.Lu + r']\w*', Keyword.Type), 104 | (r'(_[\w\']+|[' + uni.Ll + r'][\w\']*)', Name.Function), 105 | (r'--(?![!#$%&*+./<=>?@^|_~:\\]).*?$', Comment.Single), 106 | (r'\{-', Comment.Multiline, 'comment'), 107 | (r',', Punctuation), 108 | (r'[:!#$%&*+.\\/<=>?@^|~-]+', Operator), 109 | # (HACK, but it makes sense to push two instances, believe me) 110 | (r'\(', Punctuation, ('funclist', 'funclist')), 111 | (r'\)', Punctuation, '#pop:2'), 112 | ], 113 | # NOTE: the next four states are shared in the AgdaLexer; make sure 114 | # any change is compatible with Agda as well or copy over and change 115 | 'comment': [ 116 | # Multiline Comments 117 | (r'[^-{}]+', Comment.Multiline), 118 | (r'\{-', Comment.Multiline, '#push'), 119 | (r'-\}', Comment.Multiline, '#pop'), 120 | (r'[-{}]', Comment.Multiline), 121 | ], 122 | 'character': [ 123 | # Allows multi-chars, incorrectly. 124 | (r"[^\\']'", String.Char, '#pop'), 125 | (r"\\", String.Escape, 'escape'), 126 | ("'", String.Char, '#pop'), 127 | ], 128 | 'string': [ 129 | (r'[^\\"]+', String), 130 | (r"\\", String.Escape, 'escape'), 131 | ('"', String, '#pop'), 132 | ], 133 | 'escape': [ 134 | (r'[abfnrtv"\'&\\]', String.Escape, '#pop'), 135 | (r'\^[][' + uni.Lu + r'@^_]', String.Escape, '#pop'), 136 | ('|'.join(ascii), String.Escape, '#pop'), 137 | (r'o[0-7]+', String.Escape, '#pop'), 138 | (r'x[\da-fA-F]+', String.Escape, '#pop'), 139 | (r'\d+', String.Escape, '#pop'), 140 | (r'\s+\\', String.Escape, '#pop'), 141 | ], 142 | } 143 | 144 | 145 | class DAgdaLexer(RegexLexer): 146 | 147 | name = 'DAgda' 148 | aliases = ['Dagda'] 149 | filenames = ['*.agda'] 150 | mimetypes = ['text/x-agda'] 151 | 152 | reserved = ['abstract', 'codata', 'coinductive', 'constructor', 'data', 153 | 'field', 'forall', 'hiding', 'inductive', 'infix', 154 | 'infixl', 'infixr', 'instance', 'let', 'mutual', 'open', 155 | 'pattern', 'postulate', 'primitive', 'private', 156 | 'quote', 'quoteGoal', 'quoteTerm', 157 | 'record', 'renaming', 'rewrite', 'syntax', 'tactic', 158 | 'unquote', 'unquoteDecl', 'using', 'where', 'with'] 159 | 160 | tokens = { 161 | 'root': [ 162 | # Declaration 163 | (r'^(\s*)([^\s(){}]+)(\s*)(:)(\s*)', 164 | bygroups(Text, Name.Function, Text, Operator.Word, Text)), 165 | # Comments 166 | (r'--(?![!#$%&*+./<=>?@^|_~:\\]).*?$', Comment.Single), 167 | (r'\{-', Comment.Multiline, 'comment'), 168 | # Holes 169 | (r'\{!', Comment.Directive, 'hole'), 170 | # Lexemes: 171 | # Identifiers 172 | (r'\b(%s)(?!\')\b' % '|'.join(reserved), Keyword.Reserved), 173 | (r'(import|module)(\s+)', bygroups(Keyword.Reserved, Text), 'module'), 174 | (u'\\b(Set|Prop)[\u2080-\u2089]*\\b', Keyword.Type), 175 | # Special Symbols 176 | (r'(\(|\)|\{|\})', Operator), 177 | (u'(\\.{1,3}|\\||\u03BB|\u2200|\u2192|:|=|->)', Operator.Word), 178 | # Numbers 179 | (r'\d+[eE][+-]?\d+', Number.Float), 180 | (r'\d+\.\d+([eE][+-]?\d+)?', Number.Float), 181 | (r'0[xX][\da-fA-F]+', Number.Hex), 182 | (r'\d+', Number.Integer), 183 | # Strings 184 | (r"'", String.Char, 'character'), 185 | (r'"', String, 'string'), 186 | (r'[^\s(){}]+', Text), 187 | (r'\s+?', Text), # Whitespace 188 | ], 189 | 'hole': [ 190 | # Holes 191 | (r'[^!{}]+', Comment.Directive), 192 | (r'\{!', Comment.Directive, '#push'), 193 | (r'!\}', Comment.Directive, '#pop'), 194 | (r'[!{}]', Comment.Directive), 195 | ], 196 | 'module': [ 197 | (r'\{-', Comment.Multiline, 'comment'), 198 | (r'[a-zA-Z][\w.]*', Name, '#pop'), 199 | (r'[\W0-9_]+', Text) 200 | ], 201 | 'comment': DHaskellLexer.tokens['comment'], 202 | 'character': DHaskellLexer.tokens['character'], 203 | 'string': DHaskellLexer.tokens['string'], 204 | 'escape': DHaskellLexer.tokens['escape'] 205 | } 206 | -------------------------------------------------------------------------------- /Regexp/examples.agda: -------------------------------------------------------------------------------- 1 | open import Data.Product using (_×_; _,_; proj₁; proj₂) 2 | open import Data.Empty using (⊥; ⊥-elim) 3 | open import Data.Unit using (⊤; tt) 4 | open import Data.Bool using (T; Bool; true; false; if_then_else_; _∧_) 5 | open import Relation.Nullary using (Dec; ¬_; yes; no) 6 | open import Relation.Nullary.Decidable using (⌊_⌋) 7 | import Relation.Binary.PropositionalEquality as Eq 8 | open Eq using (_≡_; _≢_; refl; cong; sym; subst; trans) 9 | open import Equivalence 10 | open import Data.Nat using (ℕ; _≟_; zero; suc) 11 | 12 | module examples where 13 | 14 | a = zero 15 | b = suc zero 16 | c = suc (suc zero) 17 | 18 | Σ = ℕ 19 | open import String Σ hiding (foldl) 20 | open import Regexp Σ 21 | open import Brzozowski Σ _≟_ 22 | open import RegexpNfa Σ _≟_ as RegNFA hiding (_∈?_) 23 | 24 | a* = Atom a * 25 | b* = Atom b * 26 | 27 | ⟨aa⟩* = (Atom a · Atom a) * 28 | a*b?a* = Atom a * · (Atom b + ⟨ε⟩) · Atom a * 29 | a*b* = Atom a * · Atom b * 30 | [a[c+b]]* = (Atom a · (Atom c + Atom b)) * 31 | 32 | ⟨a+b⟩*a = (Atom a + Atom b) * · Atom a 33 | 34 | module membership-examples where 35 | x : (a ∷ a ∷ a ∷ a ∷ []) ∈ ⟨aa⟩* 36 | x = in-*2 (in-· (in-c a) (in-c a)) 37 | (in-*2 (in-· (in-c a) (in-c a)) in-*1) 38 | y : (a ∷ b ∷ a ∷ []) ∈ a*b?a* 39 | y = in-· (in-· (in-*2 (in-c a) in-*1) 40 | (in+l (in-c b))) 41 | (in-*2 (in-c a) in-*1) 42 | 43 | z : (a ∷ a ∷ b ∷ b ∷ b ∷ []) ∈ a*b* 44 | z = in-· (in-*2 (in-c a) 45 | (in-*2 (in-c a) in-*1)) 46 | (in-*2 (in-c b) 47 | (in-*2 (in-c b) 48 | (in-*2 (in-c b) in-*1))) 49 | 50 | v : (a ∷ b ∷ a ∷ c ∷ []) ∈ [a[c+b]]* 51 | v = in-*2 (in-· (in-c a) (in+r (in-c b))) 52 | (in-*2 53 | (in-· (in-c a) (in+l (in-c c))) in-*1) 54 | 55 | _ : (a ∷ b ∷ a ∷ []) ∈ ⟨a+b⟩*a 56 | _ = in-· (in-*2 (in+l (in-c a)) (in-*2 (in+r (in-c b)) in-*1)) (in-c a) 57 | 58 | open Nullable 59 | 60 | _ = {! (b ∷ []) ∈? ⟨a+b⟩*a !} 61 | 62 | 63 | module nullable-examples where 64 | x : Nullable ⟨aa⟩* 65 | x = null* 66 | 67 | y : Nullable a*b?a* 68 | y = null· (null· null* (null+r null⟨ε⟩)) null* 69 | 70 | z : Nullable a*b* 71 | z = null· null* null* 72 | 73 | u : Nullable [a[c+b]]* 74 | u = null* 75 | 76 | v : ¬ Nullable (Atom a) 77 | v = λ () 78 | 79 | module derivative-examples where 80 | x : ⟨aa⟩* [ a ] ≡ ⟨ε⟩ · Atom a · ⟨aa⟩* 81 | x = refl 82 | 83 | y : ⟨aa⟩* [ b ] ≡ ⟨⟩ · Atom a · ⟨aa⟩* 84 | y = refl 85 | 86 | yp : ∀{s} → ¬ s ∈ ⟨aa⟩* [ b ] 87 | yp (in-· (in-· () _) _) 88 | 89 | x1 : (a ∷ a ∷ a ∷ a ∷ []) ∈ ⟨aa⟩* 90 | x1 = in-*2 (in-· (in-c a) (in-c a)) 91 | (in-*2 (in-· (in-c a) (in-c a)) in-*1) 92 | 93 | x2 : (a ∷ a ∷ a ∷ []) ∈ (⟨aa⟩* [ a ]) 94 | x2 = in-· (in-· in-ε (in-c a)) 95 | (in-*2 (in-· (in-c a) (in-c a)) in-*1) 96 | 97 | x3 : (a ∷ a ∷ []) ∈ (⟨aa⟩* [ a ] [ a ]) 98 | x3 = in-· (in+r in-ε) 99 | (in-*2 (in-· (in-c a) (in-c a)) in-*1) 100 | 101 | module decidable-Brzozowski-examples where 102 | v = (a ∷ b ∷ a ∷ a ∷ a ∷ a ∷ a ∷ a ∷ a ∷ a ∷ []) 103 | 104 | e1 = v ∈? a*b?a* 105 | 106 | _ = {! (a ∷ []) ∈? ⟨aa⟩* !} 107 | 108 | e2 = (b ∷ b ∷ []) ∈? a*b?a* 109 | 110 | x : v ∈ a*b?a* 111 | x = (in-· (in-· (in-*2 (in-c a) in-*1) (in+l (in-c b))) 112 | (in-*2 (in-c a) 113 | (in-*2 (in-c a) 114 | (in-*2 (in-c a) 115 | (in-*2 (in-c a) 116 | (in-*2 (in-c a) 117 | (in-*2 (in-c a) (in-*2 (in-c a) (in-*2 (in-c a) in-*1))))))))) 118 | 119 | module dfa-examples where 120 | open import Dfa Σ 121 | open import Data.Fin renaming (_≟_ to _≟ᶠ_) 122 | open import Data.Fin.Subset 123 | open import Data.Vec hiding (foldl; map; _++_) 124 | open import Data.List hiding(_++_) 125 | 126 | TransitionsList = λ n → List (Fin n × Σ × Fin n) 127 | 128 | make-δ : ∀{n} 129 | → Fin n 130 | → TransitionsList n 131 | → (Fin n → Σ → Fin n) 132 | make-δ err [] = λ _ _ → err 133 | make-δ err ((q , x , p) ∷ xs) 134 | = λ h y → if ⌊ q ≟ᶠ h ⌋ ∧ ⌊ x ≟ y ⌋ 135 | then p 136 | else make-δ err xs h y 137 | 138 | make-dfa : (n : ℕ) 139 | → Fin n 140 | → Fin n 141 | → List (Fin n) 142 | → TransitionsList n 143 | → Dfa n 144 | make-dfa n start error finals transitions 145 | = record 146 | { S = start 147 | ; δ = make-δ error transitions 148 | ; F = ⋃ (map ⁅_⁆ finals) 149 | } 150 | 151 | dfa-⟨aa⟩* = make-dfa 3 0F 2F (0F ∷ []) ( 152 | (0F , a , 1F) 153 | ∷ (1F , a , 0F) 154 | ∷ [] 155 | ) 156 | 157 | p0 : dfa-⟨aa⟩* ↓ (a ∷ a ∷ []) 158 | p0 = tt 159 | 160 | p1 : ¬ dfa-⟨aa⟩* ↓ (a ∷ a ∷ a ∷ []) 161 | p1 () 162 | 163 | dfa-a*b?a* = make-dfa 3 0F 2F (0F ∷ 1F ∷ []) ( 164 | (0F , a , 0F) 165 | ∷ (0F , b , 1F) 166 | ∷ (1F , a , 1F) 167 | ∷ [] 168 | ) 169 | 170 | p2 : dfa-a*b?a* ↓ (a ∷ b ∷ a ∷ []) 171 | p2 = tt 172 | 173 | p3 : ¬ dfa-a*b?a* ↓ (b ∷ b ∷ a ∷ []) 174 | p3 () 175 | 176 | p4 : dfa-a*b?a* ↓ (a ∷ a ∷ a ∷ []) 177 | p4 = tt 178 | 179 | dfa-a*b* = make-dfa 3 0F 2F (0F ∷ 1F ∷ []) ( 180 | (0F , a , 0F) 181 | ∷ (0F , b , 1F) 182 | ∷ (1F , b , 1F) 183 | ∷ [] 184 | ) 185 | 186 | p5 : dfa-a*b* ↓ (a ∷ a ∷ b ∷ []) 187 | p5 = tt 188 | 189 | p6 : dfa-a*b* ↓ (b ∷ b ∷ b ∷ []) 190 | p6 = tt 191 | 192 | p7 : ¬ dfa-a*b* ↓ (c ∷ []) 193 | p7 () 194 | 195 | dfa-[a[c+b]]* = make-dfa 3 0F 2F (0F ∷ []) ( 196 | (0F , a , 1F) 197 | ∷ (1F , b , 0F) 198 | ∷ (1F , c , 0F) 199 | ∷ [] 200 | ) 201 | 202 | p8 : dfa-[a[c+b]]* ↓ [] 203 | p8 = tt 204 | 205 | p9 : dfa-[a[c+b]]* ↓ (a ∷ c ∷ a ∷ b ∷ []) 206 | p9 = tt 207 | 208 | p10 : ¬ dfa-[a[c+b]]* ↓ (a ∷ a ∷ []) 209 | p10 () 210 | 211 | open Data.Nat._≤_ 212 | 213 | 214 | 215 | dfa-binary-multiples-5 = make-dfa 6 0F 5F (0F ∷ []) ( 216 | (0F , 0 , 0F) 217 | ∷ (0F , 1 , 1F) 218 | ∷ (1F , 0 , 2F) 219 | ∷ (1F , 1 , 3F) 220 | ∷ (2F , 1 , 0F) 221 | ∷ (2F , 0 , 4F) 222 | ∷ (3F , 1 , 2F) 223 | ∷ (3F , 0 , 1F) 224 | ∷ (4F , 1 , 4F) 225 | ∷ (4F , 0 , 3F) 226 | ∷ [] 227 | ) 228 | 229 | -- [115] dec = [1110011] bin 230 | 231 | pumpLem = proj₂ 232 | (pumpingLemma dfa-binary-multiples-5) 233 | (1 ∷ 1 ∷ 1 ∷ 0 ∷ 0 ∷ 1 ∷ 1 ∷ []) 234 | tt 235 | (s≤s (s≤s (s≤s (s≤s (s≤s (s≤s (s≤s z≤n))))))) 236 | 237 | x : String 238 | x = 1 ∷ 1 ∷ [] 239 | 240 | y : String 241 | y = 1 ∷ 0 ∷ 0 ∷ [] 242 | 243 | z : String 244 | z = 1 ∷ 1 ∷ [] 245 | 246 | acc1 : dfa-binary-multiples-5 ↓ (x ++ y ^ 0 ++ z) 247 | acc1 = tt 248 | 249 | acc2 : dfa-binary-multiples-5 ↓ (x ++ y ^ 3 ++ z) 250 | acc2 = tt 251 | 252 | module nfa-examples where 253 | open import Nfa Σ 254 | open import Data.Fin renaming (_≟_ to _≟ᶠ_) 255 | open import Data.Fin.Subset 256 | open import Data.Vec hiding (foldl; map; _++_; concat) 257 | open import Data.List hiding (_++_; concat) 258 | 259 | NfaTransitionsList = λ n → List (Fin n × Σ × List (Fin n)) 260 | 261 | make-nfa-δ : ∀{n} 262 | → Fin n 263 | → NfaTransitionsList n 264 | → (Fin n → Σ → Subset n) 265 | make-nfa-δ err [] = λ _ _ → ⁅ err ⁆ 266 | make-nfa-δ err ((q , x , ps) ∷ xs) 267 | = λ h y → if ⌊ q ≟ᶠ h ⌋ ∧ ⌊ x ≟ y ⌋ 268 | then ⋃ (map ⁅_⁆ ps) 269 | else make-nfa-δ err xs h y 270 | 271 | make-nfa : (n : ℕ) 272 | → Fin n 273 | → Fin n 274 | → List (Fin n) 275 | → NfaTransitionsList n 276 | → Nfa n 277 | make-nfa n start error finals transitions 278 | = record 279 | { S = start 280 | ; δ = make-nfa-δ error transitions 281 | ; F = ⋃ (map ⁅_⁆ finals) 282 | } 283 | 284 | 285 | nfa-babb-substring = make-nfa 6 0F 5F (4F ∷ []) ( 286 | (0F , a , 0F ∷ []) 287 | ∷ (0F , b , 0F ∷ 1F ∷ []) 288 | ∷ (1F , a , 2F ∷ []) 289 | ∷ (2F , b , 3F ∷ []) 290 | ∷ (3F , b , 4F ∷ []) 291 | ∷ (4F , a , 4F ∷ []) 292 | ∷ (4F , b , 4F ∷ []) 293 | ∷ [] 294 | ) 295 | 296 | babb : String 297 | babb = b ∷ a ∷ b ∷ b ∷ [] 298 | 299 | x : nfa-babb-substring ↓ babb 300 | x = tt 301 | 302 | evaluatePath : ∀{n}{A : Nfa n} → (q : Fin n) → ∀(xs) → T(accepts A q xs) → List (Fin n) 303 | evaluatePath q [] ac = [] 304 | evaluatePath {n}{A} q (x ∷ xs) ac with nextState {n} {x} {xs} {q} {A} ac 305 | ... | p , inz , ok = p ∷ (evaluatePath {n}{A} p xs ok) 306 | 307 | y : nfa-babb-substring ↓ (a ∷ babb ++ b ∷ b ∷ []) 308 | y = tt 309 | 310 | z : ¬ nfa-babb-substring ↓ (a ∷ b ∷ a ∷ []) 311 | z () 312 | 313 | nfa-term-by-abc = make-nfa 5 0F 4F (3F ∷ []) ( 314 | (0F , a , 0F ∷ 1F ∷ []) 315 | ∷ (1F , b , 2F ∷ []) 316 | ∷ (2F , c , 3F ∷ []) 317 | ∷ [] 318 | ) 319 | 320 | abc : String 321 | abc = a ∷ b ∷ c ∷ [] 322 | 323 | t : nfa-term-by-abc ↓ (abc) 324 | t = tt 325 | 326 | u : nfa-term-by-abc ↓ (a ∷ a ∷ a ∷ abc) 327 | u = tt 328 | 329 | v : ¬ nfa-term-by-abc ↓ (b ∷ a ∷ c ∷ []) 330 | v () 331 | 332 | nfa-union-babb-abc = union nfa-babb-substring nfa-term-by-abc 333 | 334 | q : nfa-union-babb-abc ↓ (abc) 335 | q = tt 336 | 337 | r : nfa-union-babb-abc ↓ (b ∷ a ∷ b ∷ b ∷ []) 338 | r = tt 339 | 340 | s : ¬ nfa-union-babb-abc ↓ (b ∷ a ∷ c ∷ []) 341 | s () 342 | 343 | nfa-concat-babb-abc = concat nfa-babb-substring nfa-term-by-abc 344 | 345 | o : nfa-concat-babb-abc ↓ (babb ++ abc) 346 | o = tt 347 | 348 | p : ¬ nfa-concat-babb-abc ↓ babb 349 | p () 350 | 351 | nfa-star-term-abc = star nfa-term-by-abc 352 | 353 | k : nfa-star-term-abc ↓ (abc ++ abc) 354 | k = tt 355 | 356 | l : nfa-star-term-abc ↓ (abc ++ a ∷ a ∷ abc) 357 | l = tt 358 | 359 | m : ¬ nfa-star-term-abc ↓ (a ∷ abc ++ c ∷ c ∷ []) 360 | m () 361 | 362 | module regexp-nfa-examples where 363 | open import Nfa Σ 364 | open Nfa 365 | 366 | tnf = toNFA a*b?a* 367 | 368 | 369 | nfa = proj₁ (proj₂ tnf) 370 | iff = proj₂ (proj₂ tnf) 371 | 372 | -- x = (a ∷ a ∷ b ∷ a ∷ a ∷ a ∷ []) RegNFA.∈? a*b?a* 373 | -------------------------------------------------------------------------------- /Regexp/Dfa.agda: -------------------------------------------------------------------------------- 1 | open import Data.Nat as ℕ using (ℕ; zero; suc; _≥_; _≤_; _<_; s≤s; z≤n; _+_; _*_) 2 | open import Data.Nat.Properties 3 | open import Data.Fin using (Fin; toℕ; inject+; raise; 0F; 1F; 2F; 3F; 4F; 5F; 6F) renaming (_<_ to _<ᶠ_; zero to fzero; suc to fsuc) 4 | open import Data.Fin.Properties using (pigeonhole; toℕ-inject+) 5 | open import Data.Bool using (Bool; false; true; not; T) 6 | open import Data.Empty as Empty using (⊥; ⊥-elim) 7 | open import Data.Unit using (⊤; tt) 8 | open import Relation.Nullary using (Dec; yes; no; ¬_) 9 | open import Relation.Nullary.Negation using (contradiction) 10 | open import Data.Product using (_×_; Σ; ∃; ∃₂; Σ-syntax; ∃-syntax; _,_; proj₁; proj₂) 11 | open import Data.Sum using (_⊎_; inj₁; inj₂) 12 | open import Equivalence 13 | import Relation.Binary.PropositionalEquality as Eq 14 | open Eq using (_≡_; refl; _≢_; subst; sym; trans; cong; cong₂) 15 | open import Data.Vec renaming (_++_ to _++v_) hiding (take; drop) 16 | open import Data.Vec.Properties using (lookup-++ˡ) 17 | open import Data.Fin.Subset as Subset using (Subset; ∁) 18 | open import VecUtil 19 | 20 | module Dfa (Σ : Set) where 21 | open import String Σ 22 | 23 | ε : String 24 | ε = [] 25 | 26 | record Dfa (n : ℕ) : Set where 27 | field 28 | S : Fin n 29 | δ : Fin n → Σ → Fin n 30 | F : Subset n 31 | open Dfa 32 | 33 | δ^ : ∀{n} → (dfa : Dfa n) → (q : Fin n) → String → Fin n 34 | δ^ dfa q [] = q 35 | δ^ dfa q (x ∷ s) = δ^ dfa (δ dfa q x) s 36 | 37 | infix 10 _↓_ 38 | _↓_ : ∀{n} → Dfa n → String → Set 39 | dfa ↓ s = δ^ dfa (S dfa) s ∈ F dfa 40 | 41 | _↓?_ : ∀{n} → (dfa : Dfa n) → (s : String) → Dec (dfa ↓ s) 42 | dfa ↓? s = δ^ dfa (S dfa) s ∈? F dfa 43 | 44 | -------------------------------------------------------------------------------- 45 | -- Complement of a DFA 46 | 47 | ∁dfa : ∀{n} → Dfa n → Dfa n 48 | ∁dfa dfa = 49 | record 50 | { S = S dfa 51 | ; δ = δ dfa 52 | ; F = ∁ (F dfa) 53 | } 54 | 55 | ∁dfa-same-δ^ : ∀{n} → (dfa : Dfa n) (q : Fin n) (s : String) 56 | → δ^ (∁dfa dfa) q s ≡ δ^ dfa q s 57 | ∁dfa-same-δ^ dfa q [] = refl 58 | ∁dfa-same-δ^ dfa q (x ∷ s) = ∁dfa-same-δ^ dfa (δ dfa q x) s 59 | 60 | x∈p⇒x∉∁p : ∀ {n x} → (p : Subset n) → x ∈ p → x ∉ ∁ p 61 | x∈p⇒x∉∁p {_} {0F} (false ∷ p) a b = a 62 | x∈p⇒x∉∁p {_} {0F} (true ∷ p) a b = b 63 | x∈p⇒x∉∁p {_} {fsuc x} (_ ∷ p) a b = x∈p⇒x∉∁p {_}{x} p a b 64 | 65 | x∉∁p⇒x∈p : ∀ {n x} → (p : Subset n) → x ∉ ∁ p → x ∈ p 66 | x∉∁p⇒x∈p {_} {0F} (false ∷ p) ne = ne tt 67 | x∉∁p⇒x∈p {_} {0F} (true ∷ p) ne = tt 68 | x∉∁p⇒x∈p {_} {fsuc x} (_ ∷ p) ne = x∉∁p⇒x∈p {_}{x} p ne 69 | 70 | 71 | ∁dfa-correct : ∀{s n} {dfa : Dfa n} 72 | → (dfa ↓ s ⇔ ¬ (∁dfa dfa ↓ s)) 73 | ∁dfa-correct {s}{n}{dfa} rewrite ∁dfa-same-δ^ dfa (S dfa) s = 74 | (x∈p⇒x∉∁p (F dfa)) IFF (x∉∁p⇒x∈p (F dfa)) 75 | 76 | -------------------------------------------------------------------------------- 77 | -- Pumping Lemma 78 | 79 | infixl 6 _^_ 80 | _^_ : String → ℕ → String 81 | s ^ zero = [] 82 | s ^ (suc n) = s ++ s ^ n 83 | 84 | path : ∀{m} → Dfa m → Fin m → (s : String) → Vec (Fin m) (length s) 85 | path dfa q [] = [] 86 | path dfa q (c ∷ s) = q ∷ (path dfa (δ dfa q c) s) 87 | 88 | lemmaPath : ∀{m} 89 | → (dfa : Dfa m) 90 | → (s : String) 91 | → (i : Fin (length s)) 92 | → (q : Fin m) 93 | → path dfa q s ! i ≡ δ^ dfa q (take (toℕ i) s) 94 | lemmaPath dfa (c ∷ s) fzero q = refl 95 | lemmaPath dfa (c ∷ s) (fsuc i) q = lemmaPath dfa s i (δ dfa q c) 96 | 97 | lemma-δ^ : ∀{n} 98 | (dfa : Dfa n) 99 | → (s : String) 100 | → (t : String) 101 | → (q : Fin n) 102 | → δ^ dfa q (s ++ t) ≡ δ^ dfa (δ^ dfa q s) t 103 | lemma-δ^ dfa [] t q = refl 104 | lemma-δ^ dfa (c ∷ s) t q = lemma-δ^ dfa s t (δ dfa q c) 105 | 106 | returns-back : ∀{n} 107 | → (dfa : Dfa n) 108 | → (s : String) 109 | → (q : Fin n) 110 | → q ≡ δ^ dfa q s 111 | → ∀(m : ℕ) → q ≡ δ^ dfa q (s ^ m) 112 | returns-back dfa s q eq zero = refl 113 | returns-back dfa s q eq (suc m) with 114 | returns-back dfa s q eq m | lemma-δ^ dfa s (s ^ m) q 115 | ... | ind | lm2 rewrite sym eq = trans ind (sym lm2) 116 | 117 | pumping-same-state : ∀{n} {dfa : Dfa n} 118 | → (s : String) 119 | → (t : String) 120 | → (u : String) 121 | → let p = δ^ dfa (S dfa) s in 122 | p ≡ δ^ dfa p t 123 | → dfa ↓ (s ++ t ++ u) 124 | → ∀ (m : ℕ) → dfa ↓ (s ++ t ^ m ++ u) 125 | pumping-same-state {n} {dfa} s t u p≡δ^_pt dfa↓stu m with 126 | returns-back dfa t (δ^ dfa (S dfa) s) p≡δ^_pt m 127 | ... | pump with lemma-δ^ dfa (s ++ t) u (S dfa) 128 | | lemma-δ^ dfa s (t ^ m) (S dfa) 129 | | lemma-δ^ dfa s t (S dfa) 130 | | lemma-δ^ dfa (s ++ (t ^ m)) u (S dfa) 131 | ... | d1 | d2 | d3 | d4 rewrite 132 | trans pump (sym d2) 133 | | sym (trans p≡δ^_pt (sym d3)) 134 | | ++-assoc s t u 135 | | ++-assoc s (t ^ m) u 136 | | trans d1 (sym d4) = dfa↓stu 137 | 138 | take++drop : (n : ℕ) → (s : String) → s ≡ take n s ++ drop n s 139 | take++drop zero s = refl 140 | take++drop (suc n) [] = refl 141 | take++drop (suc n) (c ∷ s) = cong (c ∷_) (take++drop n s) 142 | 143 | lemminoTakeDrop1 : ∀{s : String} → {i j : Fin (length s)} 144 | → i <ᶠ j 145 | → s ≡ take (toℕ i) s ++ drop (toℕ i) (take (toℕ j) s) ++ drop (toℕ j) s 146 | lemminoTakeDrop1 {c ∷ s} {fzero} {fsuc j} (s≤s z≤n) = cong (c ∷_) (take++drop (toℕ j) s) 147 | lemminoTakeDrop1 {c ∷ s} {fsuc i} {fsuc j} (s≤s i0 : ∀{n} → (dfa : Dfa n) → 1 ≤ n 115 | dfa-states>0 {suc n} dfa = s≤s z≤n 116 | 117 | lemmaℕ≤ : (m : ℕ) → 1 ≤ m → 1 + m ≤ m + m 118 | lemmaℕ≤ m lq = +-monoˡ-≤ m lq 119 | 120 | length-++ : (xs ys : String) → length xs + length ys ≡ length (xs ++ ys) 121 | length-++ [] ys = refl 122 | length-++ (x ∷ xs) ys = cong suc (length-++ xs ys) 123 | 124 | absurd1 : ∀{s} {ts : String} → ¬ (s ∷ ts ≡ []) 125 | absurd1 = λ () 126 | 127 | listlem0 : {a b : Σ} {s t : String} → (a ∷ s) ≡ (b ∷ t) → a ≡ b 128 | listlem0 {a} {.a} {s} {.s} refl = refl 129 | 130 | listlem1 : {c : Σ} {s t : String} → (c ∷ s) ≡ (c ∷ t) → s ≡ t 131 | listlem1 refl = refl 132 | 133 | listlem2 : ∀{s t l : String} → (s ++ t) ≡ (s ++ l) → t ≡ l 134 | listlem2 {[]} eq = eq 135 | listlem2 {x ∷ s} {t} {l} eq = listlem2 (listlem1 eq) 136 | 137 | char-pow-++ : ∀{c n x y} 138 | → ((c ∷ []) ^ n) ≡ x ++ y 139 | → ∃[ l ] ∃[ m ] (x ≡ (c ∷ []) ^ l × y ≡ (c ∷ []) ^ m × n ≡ l + m) 140 | char-pow-++ {c} {0F} {[]} {[]} eq = 0 , 0 , refl , refl , refl 141 | char-pow-++ {c} {n} {[]} refl = 0 , n , refl , refl , refl 142 | char-pow-++ {c} {suc n} {x ∷ xs} {y} eq with listlem0 eq 143 | char-pow-++ {c} {suc n} {.c ∷ xs} {y} eq | refl with char-pow-++ {c} {n}{xs}{y} (listlem1 eq) 144 | ... | l , m , eq1 , eq2 , eq3 = 145 | suc l , m , cong (c ∷_) eq1 , eq2 , (cong suc eq3) 146 | 147 | power>0 : ∀{l y} → y ≡ I ^ l → y ≢ ε → 1 ≤ l 148 | power>0 {0F} refl y = ⊥-elim(y refl) 149 | power>0 {suc l} eq neq = s≤s z≤n 150 | 151 | xyz-to-power-base : ∀{n m x y z} 152 | → (I ^ n) ++ (O ^ m) ≡ x ++ y ++ z 153 | → length (x ++ y) ≤ n 154 | → ∃[ l ] ∃[ p ] ∃[ q ] (x ≡ I ^ l 155 | × y ≡ I ^ p 156 | × z ≡ (I ^ q) ++ (O ^ m)) 157 | xyz-to-power-base {0F} {0F} {[]} {[]} {[]} o t = zero , zero , zero , refl , refl , refl 158 | xyz-to-power-base {0F} {suc m} {[]} {[]} {x ∷ z} e t = zero , zero , zero , refl , refl , sym e 159 | xyz-to-power-base {suc n} {0F} {x} {y} {z} e t rewrite ++-idʳ (I ^ n) with char-pow-++ {1} {suc n} {x} {y ++ z} e 160 | ... | l , m , eq1 , eq2 , eq3 with char-pow-++ {1} {m} {y} (sym eq2) 161 | ... | l2 , m2 , eq4 , eq5 , eq6 rewrite sym (++-idʳ (I ^ m2)) = l , l2 , m2 , eq1 , eq4 , eq5 162 | xyz-to-power-base {suc n} {suc m} {[]} {[]} {x ∷ z} e t with listlem0 e 163 | xyz-to-power-base {suc n} {suc m} {[]} {[]} {.1 ∷ z} e les | refl with xyz-to-power-base {n} {suc m} {[]} {[]} {z} (listlem1 e) z≤n 164 | ... | l , p , q , eq1 , eq2 , eq3 = l , p , suc q , eq1 , eq2 , cong (1 ∷_) eq3 165 | xyz-to-power-base {suc n} {suc m} {[]} {x ∷ y} {z} e t with listlem0 e 166 | xyz-to-power-base {suc n} {suc m} {[]} {.1 ∷ y} {z} e (s≤s les) | refl with xyz-to-power-base {n} {suc m} {[]} {y} {z} (listlem1 e) les 167 | ... | l , p , q , eq1 , eq2 , eq3 = l , suc p , q , eq1 , cong (1 ∷_) eq2 , eq3 168 | xyz-to-power-base {suc n} {suc m} {c ∷ x} {y} {z} e t with listlem0 e 169 | xyz-to-power-base {suc n} {suc m} {.1 ∷ x} {y} {z} e (s≤s les) | refl with xyz-to-power-base {n} {suc m} {x} {y} {z} (listlem1 e) les 170 | ... | l , p , q , eq1 , eq2 , eq3 = suc l , p , q , cong (1 ∷_) eq1 , eq2 , eq3 171 | 172 | xyz-to-power : ∀{n x y z} 173 | → (I ^ n) ++ (O ^ n) ≡ x ++ y ++ z 174 | → length (x ++ y) ≤ n 175 | → y ≢ ε 176 | → ∃[ l ] ∃[ p ] ∃[ q ] (x ≡ I ^ l 177 | × y ≡ I ^ p 178 | × 0 < p 179 | × z ≡ (I ^ q) ++ (O ^ n)) 180 | xyz-to-power {n} {x} {y} {z} e les neq with xyz-to-power-base {n} {n} {x} {y} {z} e les 181 | ... | l , p , q , eq1 , eq2 , eq3 = l , p , q , eq1 , eq2 , power>0 eq2 neq , eq3 182 | 183 | powSame : ∀{s n m a} → (a ∷ s) ^ n ≡ (a ∷ s) ^ m → n ≡ m 184 | powSame {s} {0F} {0F} e = refl 185 | powSame {s} {0F} {suc n} () 186 | powSame {s} {suc n} {0F} () 187 | powSame {s} {suc n} {suc m} {a} eq with listlem2 {a ∷ s} {(a ∷ s) ^ n} {(a ∷ s) ^ m} eq 188 | ... | u = cong suc (powSame u) 189 | 190 | exponents-equal-base : (m p n q : ℕ) 191 | → I ^ m ++ O ^ n ≡ I ^ p ++ O ^ q 192 | → m ≡ p × n ≡ q 193 | exponents-equal-base 0F 0F _ _ eq = refl , powSame eq 194 | exponents-equal-base 0F (suc p) n q eq with char-pow-++ {0} {n} {I ^ (suc p)} {O ^ q} eq 195 | ... | zero , _ , () , _ , _ 196 | ... | suc a , _ , e , _ , _ = contradiction (listlem0 e) λ () 197 | exponents-equal-base (suc m) 0F n q eq with char-pow-++ {0} {q} {I ^ (suc m)} {O ^ n} (sym eq) 198 | ... | zero , _ , () , _ , _ 199 | ... | suc a , _ , e , _ , _ = contradiction (listlem0 e) λ () 200 | exponents-equal-base (suc m) (suc p) n q eq with exponents-equal-base m p n q (listlem1 eq) 201 | ... | fst , snd = cong suc fst , snd 202 | 203 | exponents-equal : ∀{n m} → (I ^ n ++ O ^ m) ∈L → n ≡ m 204 | exponents-equal {n} {m} (p , snd) with exponents-equal-base n p m p snd 205 | ... | m≡p , n≡p rewrite m≡p | n≡p = refl 206 | 207 | p≮p : ∀{p} → ¬ (p < p) 208 | p≮p (s≤s pp) = p≮p pp 209 | 210 | <-implies-≢ : ∀{p n} → p < n → n ≢ p 211 | <-implies-≢ pn refl = p≮p pn 212 | 213 | +-monoR-≤ : ∀{m n} (p : ℕ) → m ≤ n → m ≤ p + n 214 | +-monoR-≤ {m} {n} p mn with +-mono-≤ {_}{_}{_}{p} mn z≤n 215 | ... | u rewrite +-identityʳ m | +-comm p n = u 216 | 217 | nLess0 : (p n : ℕ) → 1 ≤ p → n < p * suc n 218 | nLess0 (suc p) zero lt = s≤s z≤n 219 | nLess0 (suc p) (suc n) lt with nLess0 (suc p) n lt 220 | ... | z with +-monoʳ-≤ 1 z 221 | ... | u rewrite *-comm p (suc (suc n)) 222 | | *-comm p (suc n) 223 | | sym (+-assoc n p (p + n * p)) 224 | | +-comm n p 225 | | +-assoc p n (p + n * p) 226 | | sym (+-suc p (n + (p + n * p))) 227 | | sym (+-suc p (suc (n + (p + n * p)))) = +-monoR-≤ p u 228 | 229 | nLess1 : (l p n q : ℕ) → 1 ≤ p → n < l + p * suc n + q 230 | nLess1 l p n q x rewrite 231 | +-assoc l (p * suc n) q 232 | | +-comm (p * suc n) q 233 | | sym (+-assoc l q (p * suc n)) = +-monoR-≤ (l + q) (nLess0 p n x) 234 | 235 | absurd-sum : (l p n q : ℕ) 236 | → 1 ≤ p 237 | → l + p * (1 + n) + q ≢ n 238 | absurd-sum l p n q lt = <-implies-≢ (nLess1 l p n q lt) 239 | 240 | L-not-regular : ¬ ∃₂ λ (n : ℕ) (dfa : Dfa n) 241 | → ∀ (s : String) 242 | → s ∈L ⇔ dfa ↓ s 243 | L-not-regular (n , dfa , s∈L⇔dfa↓s) with proj₂ 244 | (pumpingLemma dfa) 245 | (L n) 246 | (_⇔_.to (s∈L⇔dfa↓s (L n)) (n , refl)) 247 | (subst (suc n ≤_) 248 | (sym (L-length=n+n n)) 249 | (lemmaℕ≤ n (dfa-states>0 dfa)) 250 | ) 251 | ... | x , y , z , eq , neq , lm , pump with xyz-to-power eq lm neq 252 | ... | l , p , q , x≡I^l , y≡I^p , 0