├── Closed ├── Makefile ├── Structures.agda ├── Properties.agda ├── Construct.agda ├── Bundles.agda └── Instances.agda ├── Algebra └── Construct │ ├── Makefile │ └── Function.agda ├── ld.agda-lib ├── .gitignore ├── icfp21-submitted.pdf ├── Closed.agda ├── commands.tex ├── Examples.lagda ├── Reflections.lagda ├── Makefile ├── Misc.lagda ├── readme.md ├── Language.lagda ├── Symbolic.lagda ├── Transport.lagda ├── macros.tex ├── Decidability.lagda ├── Predicate ├── Algebra.agda └── Properties.lagda ├── SizedAutomatic.lagda ├── Automatic.lagda ├── final-review-responses.md ├── Predicate.lagda ├── To do.md ├── talk.tex ├── unicode.tex ├── talk.md ├── Calculus.lagda ├── Existential.lagda ├── Related work.md ├── Inverses.lagda ├── reviews.md └── response-to-reviews.md /Closed/Makefile: -------------------------------------------------------------------------------- 1 | all: 2 | make -C .. 3 | -------------------------------------------------------------------------------- /Algebra/Construct/Makefile: -------------------------------------------------------------------------------- 1 | all: 2 | make -C ../.. 3 | -------------------------------------------------------------------------------- /ld.agda-lib: -------------------------------------------------------------------------------- 1 | name: language-derivatives 2 | depend: standard-library 3 | include: . 4 | -------------------------------------------------------------------------------- /.gitignore: -------------------------------------------------------------------------------- 1 | latex/ 2 | *.agdai 3 | 4 | .paper-token 5 | .talk-token 6 | source.zip 7 | paper.zip 8 | -------------------------------------------------------------------------------- /icfp21-submitted.pdf: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/conal/paper-2021-language-derivatives/HEAD/icfp21-submitted.pdf -------------------------------------------------------------------------------- /Closed.agda: -------------------------------------------------------------------------------- 1 | -- Closed semirings 2 | {-# OPTIONS --safe --without-K #-} 3 | 4 | module Closed where 5 | 6 | open import Closed.Structures public 7 | open import Closed.Bundles public 8 | open import Closed.Instances public 9 | open import Closed.Construct public 10 | -------------------------------------------------------------------------------- /commands.tex: -------------------------------------------------------------------------------- 1 | %%%%%%%%%% AGDA ALIASES 2 | 3 | \newcommand{\AP}{\AgdaPrimitive} 4 | \newcommand{\APT}{\AgdaPrimitiveType} 5 | \newcommand{\AK}{\AgdaKeyword} 6 | \newcommand{\AM}{\AgdaModule} 7 | \newcommand{\AS}{\AgdaSymbol} 8 | \newcommand{\AStr}{\AgdaString} 9 | \newcommand{\AN}{\AgdaNumber} 10 | \newcommand{\AD}{\AgdaDatatype} 11 | \newcommand{\AF}{\AgdaFunction} 12 | \newcommand{\ARe}{\AgdaRecord} 13 | \newcommand{\AFi}{\AgdaField} 14 | \newcommand{\AB}{\AgdaBound} 15 | \newcommand{\AIC}{\AgdaInductiveConstructor} 16 | \newcommand{\AC}{\AgdaComment} 17 | \newcommand{\AG}{\AgdaGeneralizable} 18 | \newcommand{\APo}{\AgdaPostulate} 19 | 20 | \newcommand{\Set}{\mathbf{Set}} 21 | 22 | \newcommand{\AR}{\AgdaRef} 23 | 24 | %% Targets I've turned off 25 | \newcommand\NoAgdaTarget[1]{} 26 | -------------------------------------------------------------------------------- /Closed/Structures.agda: -------------------------------------------------------------------------------- 1 | -- Star Semirings Structures 2 | {-# OPTIONS --without-K --safe #-} 3 | 4 | open import Relation.Binary using (Rel) 5 | module Closed.Structures {a ℓ} {A : Set a} (_≈_ : Rel A ℓ) where 6 | 7 | open import Level 8 | open import Algebra.Core 9 | open import Algebra.Structures _≈_ 10 | 11 | Star : (_+_ _✲_ : Op₂ A) (𝟘 𝟙 : A) (_✯ : Op₁ A) → Set (a ⊔ ℓ) 12 | Star _+_ _✲_ 𝟘 𝟙 _✯ = ∀ x → (x ✯) ≈ (𝟙 + (x ✲ (x ✯))) 13 | 14 | record IsClosedSemiring (_+_ _✲_ : Op₂ A) (𝟘 𝟙 : A) (_✯ : Op₁ A) : Set (a ⊔ ℓ) where 15 | field 16 | isSemiring : IsSemiring _+_ _✲_ 𝟘 𝟙 17 | star : Star _+_ _✲_ 𝟘 𝟙 _✯ 18 | open IsSemiring isSemiring public 19 | 20 | record IsClosedCommutativeSemiring (_+_ _✲_ : Op₂ A) (𝟘 𝟙 : A) (_✯ : Op₁ A) : Set (a ⊔ ℓ) where 21 | field 22 | isCommutativeSemiring : IsCommutativeSemiring _+_ _✲_ 𝟘 𝟙 23 | star : Star _+_ _✲_ 𝟘 𝟙 _✯ 24 | 25 | open IsCommutativeSemiring isCommutativeSemiring public 26 | 27 | isClosedSemiring : IsClosedSemiring _+_ _✲_ 𝟘 𝟙 _✯ 28 | isClosedSemiring = record 29 | { isSemiring = isSemiring 30 | ; star = star 31 | } 32 | -------------------------------------------------------------------------------- /Examples.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | module Examples where 3 | 4 | open import Data.Product 5 | open import Data.Sum hiding ([_,_]) 6 | open import Data.Char 7 | open import Data.List 8 | open import Relation.Binary.PropositionalEquality hiding ([_]) 9 | open import Data.List.Relation.Unary.All 10 | 11 | open import Language Char 12 | 13 | -- pattern [_,_] x y = x ∷ y ∷ [] 14 | -- pattern [_,_,_] x y z = x ∷ y ∷ z ∷ [] 15 | -- pattern [_,_,_,_] x y z w = x ∷ y ∷ z ∷ w ∷ [] 16 | 17 | \end{code} 18 | 19 | %<*examples> 20 | \begin{center} 21 | \begin{minipage}[t]{20ex} 22 | \begin{code} 23 | a∪b : Lang 24 | a∪b = ` 'a' ∪ ` 'b' 25 | 26 | a⋆b : Lang 27 | a⋆b = ` 'a' ⋆ ` 'b' 28 | 29 | a∪b☆ : Lang 30 | a∪b☆ = a∪b ☆ 31 | \end{code} 32 | \end{minipage} 33 | \hspace{8ex} 34 | \begin{minipage}[t]{40ex} 35 | \begin{code} 36 | _ : a∪b [ 'b' ] 37 | _ = inj₂ refl 38 | 39 | _ : a⋆b ('a' ∷ 'b' ∷ []) 40 | _ = ([ 'a' ] , [ 'b' ]) , refl , refl , refl 41 | 42 | _ : a∪b☆ ('a' ∷ 'b' ∷ 'a' ∷ []) 43 | _ = [ 'a' ] ∷ [ 'b' ] ∷ [ 'a' ] ∷ [] 44 | , refl 45 | , inj₁ refl ∷ inj₂ refl ∷ inj₁ refl ∷ [] 46 | \end{code} 47 | \end{minipage} 48 | \end{center} 49 | % 50 | -------------------------------------------------------------------------------- /Closed/Properties.agda: -------------------------------------------------------------------------------- 1 | -- Star semirings properties 2 | 3 | {-# OPTIONS --safe #-} 4 | {-# OPTIONS --without-K #-} 5 | 6 | open import Closed.Bundles 7 | 8 | module Closed.Properties {c ℓ} (R : ClosedSemiring c ℓ) where 9 | 10 | -- If I want to prove properties of closed *commutative* semirings, then move 11 | -- the R parameter to a local module, and add another. 12 | 13 | open import Function using (_↔_) 14 | 15 | -- open import Isomorphisms using (module ↔R) 16 | open import Closed.Structures 17 | 18 | open ClosedSemiring R renaming (Carrier to A) 19 | 20 | open import Relation.Binary.Reasoning.Setoid setoid 21 | 22 | -- if `x ≈ a * x + b` then `x ≈ a ✯ * b`. 23 | 24 | private 25 | variable 26 | x a b : A 27 | 28 | _isFixpointOf_ : A → (A → A) → Set _ 29 | a isFixpointOf f = a ≈ f a 30 | 31 | isFix : (a ✯ * b) isFixpointOf (λ x → b + a * x) 32 | isFix {a}{b} = 33 | begin 34 | a ✯ * b 35 | ≈⟨ *-congʳ (star _) ⟩ 36 | (1# + a * a ✯) * b 37 | ≈⟨ distribʳ _ _ _ ⟩ 38 | 1# * b + (a * a ✯) * b 39 | ≈⟨ +-cong (*-identityˡ _) (*-assoc _ _ _) ⟩ 40 | b + a * (a ✯ * b) 41 | ∎ 42 | 43 | -- TODO: Is there a general form of Arden's rule for arbitrary closed semirings? 44 | -- This rule says that `a ✯ * b` is the *least* fixpoint of `λ x → b + a * x`. 45 | -------------------------------------------------------------------------------- /Closed/Construct.agda: -------------------------------------------------------------------------------- 1 | -- Closed semiring constructions 2 | {-# OPTIONS --safe --without-K #-} 3 | 4 | module Closed.Construct where 5 | 6 | open import Level 7 | open import Algebra 8 | 9 | open import Closed.Structures 10 | open import Closed.Bundles 11 | open import Existential ; open Inj 12 | 13 | private variable b c ℓ : Level 14 | 15 | module _ (r : ClosedSemiring c ℓ) (f : ClosedSemiring.Carrier r → Set b) where 16 | open ClosedSemiring r hiding (isSemiring) ; open Dop f 17 | mkClosedSemiring : D₂ _+_ → D₂ _*_ → D₀ 0# → D₀ 1# → D₁ _✯ 18 | → ClosedSemiring (c ⊔ b) ℓ 19 | mkClosedSemiring _+′_ _*′_ 0#′ 1#′ _✯′ = record 20 | { _✯ = inj₁ _✯′ 21 | ; isClosedSemiring = record 22 | { isSemiring = isSemiring 23 | ; star = prop₁ star 24 | } 25 | } where open Semiring (mkSemiring semiring f _+′_ _*′_ 0#′ 1#′) 26 | 27 | module _ (r : ClosedCommutativeSemiring c ℓ) 28 | (f : ClosedCommutativeSemiring.Carrier r → Set b) where 29 | open ClosedCommutativeSemiring r hiding (isCommutativeSemiring) ; open Dop f 30 | mkClosedCommutativeSemiring : D₂ _+_ → D₂ _*_ → D₀ 0# → D₀ 1# → D₁ _✯ 31 | → ClosedCommutativeSemiring (c ⊔ b) ℓ 32 | mkClosedCommutativeSemiring _+′_ _*′_ 0#′ 1#′ _✯′ = record 33 | { _✯ = inj₁ _✯′ 34 | ; isClosedCommutativeSemiring = record 35 | { isCommutativeSemiring = isCommutativeSemiring 36 | ; star = prop₁ star 37 | } 38 | } where open CommutativeSemiring 39 | (mkCommutativeSemiring commutativeSemiring f _+′_ _*′_ 0#′ 1#′) 40 | -------------------------------------------------------------------------------- /Reflections.lagda: -------------------------------------------------------------------------------- 1 | \begin{code}[hide] 2 | module Reflections {ℓ} (A : Set ℓ) where 3 | \end{code} 4 | 5 | \begin{comment} 6 | \begin{frame}{Where are we?} 7 | \begin{itemize}\itemsep3ex 8 | \item Simple formulation of languages with membership proofs (parsings). 9 | %% \item Algebraic properties from standard constructions. 10 | \item Recursive decomposition lemmas. 11 | \item Recursive decidability lemmas. 12 | \end{itemize} 13 | \vspace{4ex} 14 | Next: Assemble lemmas into implementations. 15 | \end{frame} 16 | \end{comment} 17 | 18 | \begin{code}[hide] 19 | open import Level 20 | open import Function using (_∘_;_↔_) 21 | open import Relation.Binary.PropositionalEquality using (_≗_) 22 | open import Data.Product using (_×_) 23 | open import Data.List using (foldl) 24 | 25 | open import Misc {ℓ} 26 | open import Language A 27 | open import Calculus A 28 | open import Decidability {ℓ} 29 | open import Inverses {ℓ} 30 | 31 | private 32 | variable 33 | a : A 34 | b : Level 35 | B : Set b 36 | P Q : Lang 37 | 38 | private 39 | module Stuff where 40 | \end{code} 41 | 42 | %<*summary-a> 43 | \begin{code} 44 | ν⇂ : Lang → Set ℓ 45 | δ⇂ : Lang → A → Lang 46 | 47 | ν∘foldlδ⇂ : ν ∘ foldl δ P ≗ P 48 | \end{code} 49 | \begin{code}[hide] 50 | ν⇂ = ν⇃ 51 | δ⇂ = δ⇃ 52 | ν∘foldlδ⇂ = ν∘foldlδ 53 | \end{code} 54 | % 55 | 56 | %<*summary-b> 57 | \begin{code} 58 | ν⋆⇂ : ν (P ⋆ Q) ↔ (ν P × ν Q) 59 | δ⋆⇂ : δ (P ⋆ Q) a ⟷ ν P · δ Q a ∪ δ P a ⋆ Q 60 | -- etc 61 | \end{code} 62 | \begin{code}[hide] 63 | ν⋆⇂ = ν⋆ 64 | δ⋆⇂ = δ⋆ 65 | \end{code} 66 | % 67 | -------------------------------------------------------------------------------- /Closed/Bundles.agda: -------------------------------------------------------------------------------- 1 | -- Star semirings bundles 2 | 3 | {-# OPTIONS --without-K --safe #-} 4 | 5 | open import Algebra.Bundles 6 | 7 | module Closed.Bundles where 8 | 9 | open import Level 10 | open import Relation.Binary using (Rel) 11 | open import Algebra.Core 12 | 13 | open import Closed.Structures 14 | 15 | record ClosedSemiring c ℓ : Set (suc (c ⊔ ℓ)) where 16 | infixl 10 _✯ 17 | infixl 7 _*_ 18 | infixl 6 _+_ 19 | infix 4 _≈_ 20 | field 21 | Carrier : Set c 22 | _≈_ : Rel Carrier ℓ 23 | _+_ : Op₂ Carrier 24 | _*_ : Op₂ Carrier 25 | 0# : Carrier 26 | 1# : Carrier 27 | _✯ : Op₁ Carrier 28 | isClosedSemiring : IsClosedSemiring _≈_ _+_ _*_ 0# 1# _✯ 29 | 30 | open IsClosedSemiring isClosedSemiring public 31 | 32 | semiring : Semiring _ _ 33 | semiring = record { isSemiring = isSemiring } 34 | 35 | record ClosedCommutativeSemiring c ℓ : Set (suc (c ⊔ ℓ)) where 36 | infixl 10 _✯ 37 | infixl 7 _*_ 38 | infixl 6 _+_ 39 | infix 4 _≈_ 40 | field 41 | Carrier : Set c 42 | _≈_ : Rel Carrier ℓ 43 | _+_ : Op₂ Carrier 44 | _*_ : Op₂ Carrier 45 | 0# : Carrier 46 | 1# : Carrier 47 | _✯ : Op₁ Carrier 48 | isClosedCommutativeSemiring : IsClosedCommutativeSemiring _≈_ _+_ _*_ 0# 1# _✯ 49 | 50 | open IsClosedCommutativeSemiring isClosedCommutativeSemiring public 51 | 52 | commutativeSemiring : CommutativeSemiring _ _ 53 | commutativeSemiring = record { isCommutativeSemiring = isCommutativeSemiring } 54 | 55 | closedSemiring : ClosedSemiring _ _ 56 | closedSemiring = record { isClosedSemiring = isClosedSemiring } 57 | -------------------------------------------------------------------------------- /Makefile: -------------------------------------------------------------------------------- 1 | PAPER=paper 2 | TALK=talk 3 | 4 | # all: latex/$(PAPER).pdf 5 | all: latex/$(TALK).pdf 6 | 7 | MODULES:= \ 8 | Language \ 9 | Examples \ 10 | Inverses \ 11 | Calculus \ 12 | Decidability \ 13 | Reflections \ 14 | Symbolic \ 15 | Automatic \ 16 | SizedAutomatic \ 17 | Predicate \ 18 | Existential \ 19 | Transport 20 | 21 | LAGDAS:=$(patsubst %,%.lagda,$(MODULES)) 22 | 23 | AGDA_DEPENDENCIES:=$(patsubst %,latex/%.tex,$(MODULES)) 24 | .SECONDARY: $(AGDA_DEPENDENCIES) 25 | 26 | LATEX_DEPENDENCIES:= \ 27 | latex/bib.bib \ 28 | latex/acm-bib.bib \ 29 | latex/macros.tex \ 30 | latex/unicode.tex \ 31 | latex/commands.tex \ 32 | $(AGDA_DEPENDENCIES) 33 | 34 | test : 35 | echo $(LATEX_DEPENDENCIES) 36 | 37 | AGDA=agda 38 | 39 | # AGDA-EXTRAS=--only-scope-checking 40 | 41 | PRECIOUS: $(LATEX_DEPENDENCIES) latex/$(PAPER).tex latex/$(TALK).tex 42 | 43 | latex/%.tex: %.lagda 44 | @mkdir -p $(dir $@) 45 | ${AGDA} -i . --latex --latex-dir=latex $(AGDA-EXTRAS) $< 46 | 47 | # > $(basename $@).log 48 | 49 | latex/%: % 50 | @mkdir -p $(dir $@) 51 | cp $< $@ 52 | 53 | latex/%.pdf: $(LATEX_DEPENDENCIES) latex/%.tex 54 | cd latex && latexmk -xelatex -bibtex $*.tex 55 | @touch $@ 56 | 57 | # The touch is in case latexmk decides not to update the pdf. 58 | 59 | SHOWPDF=skim 60 | 61 | see: $(PAPER).see 62 | 63 | %.see: latex/%.pdf 64 | ${SHOWPDF} $< 65 | 66 | SOURCES=$(shell find . -name '*.*agda' | grep -v Junk | grep -v _build) 67 | 68 | code.zip: $(SOURCES) ld.agda-lib 69 | zip $@ $^ 70 | 71 | ACMART=/usr/local/texlive/2021/texmf-dist/tex/latex/acmart 72 | 73 | paper.zip: latex/*.tex latex/*.bbl latex/*.sty $(ACMART)/acmart.cls 74 | zip $@ $^ 75 | 76 | clean: 77 | rm -r latex 78 | 79 | tags: $(SOURCES) paper.tex talk.tex 80 | etags $^ 81 | 82 | web: .paper-token .talk-token 83 | 84 | .paper-token: latex/$(PAPER).pdf 85 | scp $< conal@conal.net:/home/conal/domains/conal/htdocs/papers/language-derivatives/ 86 | @touch $@ 87 | 88 | .talk-token: latex/$(TALK).pdf 89 | scp $< conal@conal.net:/home/conal/domains/conal/htdocs/talks/language-derivatives.pdf 90 | @touch $@ 91 | 92 | -------------------------------------------------------------------------------- /Misc.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | {-# OPTIONS --safe --without-K #-} 3 | 4 | open import Level 5 | 6 | module Misc {ℓ : Level} where 7 | 8 | \end{code} 9 | 10 | \subsection{̄{⊥ and ⊤} 11 | 12 | Make ⊥ and ⊤ level-monomorphic to simplify signatures elsewhere. 13 | Doing so loses some generality, but that generality is already lost in equational reasoning on types. 14 | 15 | \begin{code} 16 | open import Data.Empty.Polymorphic using () renaming (⊥ to ⊥′) 17 | open import Data.Empty.Polymorphic using (⊥-elim) public 18 | open import Data.Unit.Polymorphic using () renaming (⊤ to ⊤′) 19 | 20 | ⊥ : Set ℓ 21 | ⊥ = ⊥′ 22 | 23 | ⊤ : Set ℓ 24 | ⊤ = ⊤′ 25 | 26 | import Data.Unit 27 | pattern tt = lift Data.Unit.tt 28 | 29 | infix 3 ¬_ 30 | ¬_ : Set ℓ → Set ℓ 31 | ¬ X = X → ⊥ 32 | 33 | open import Data.Unit using () renaming (tt to tt⇂) public 34 | open import Data.Empty using () renaming (⊥-elim to ⊥-elim⇂) public 35 | \end{code} 36 | 37 | \subsection{Lists} 38 | 39 | For succinctness and consistency with algebraically related notions, use star for list types. 40 | So that I can tweak the typesetting, rename append. 41 | 42 | \begin{code} 43 | open import Data.List 44 | 45 | infixl 10 _✶ 46 | _✶ : Set ℓ → Set ℓ 47 | _✶ = List 48 | 49 | infixr 5 _⊙_ 50 | _⊙_ : ∀ {A} → A ✶ → A ✶ → A ✶ 51 | _⊙_ = _++_ 52 | \end{code} 53 | 54 | \subsection{Functions} 55 | 56 | \begin{code} 57 | open import Data.Product 58 | 59 | uncurry₃ˡ : ∀ {a b c d} {A : Set a} {B : A → Set b} {C : (x : A) → B x → Set c} 60 | {D : Σ (Σ A B) (uncurry C) → Set d} → 61 | ((x : A) → (y : B x) → (z : C x y) → D ((x , y) , z)) → 62 | ((p : _) → D p) 63 | uncurry₃ˡ f ((x , y) , z) = f x y z 64 | 65 | \end{code} 66 | 67 | %% %<*Ops> 68 | %% \AgdaTarget{Op₁, Op₂} 69 | %% \begin{code} 70 | %% Op₁ : ∀ {ℓ} → Set ℓ → Set ℓ 71 | %% Op₁ A = A → A 72 | 73 | %% Op₂ : ∀ {ℓ} → Set ℓ → Set ℓ 74 | %% Op₂ A = A → A → A 75 | %% \end{code} 76 | %% % 77 | 78 | \subsection{Quantification} 79 | 80 | \begin{code} 81 | open import Data.Product 82 | private variable A B : Set ℓ 83 | \end{code} 84 | 85 | Agda doesn't like ``∃ λ (a , b) → ...'', maybe because it can't sort out how \AB{b} could depends on \AB{a}. 86 | I'd like to find a better solution. 87 | \begin{code} 88 | ∃⇃ : (A × B → Set ℓ) → Set ℓ 89 | ∃⇃ = ∃ 90 | \end{code} 91 | -------------------------------------------------------------------------------- /readme.md: -------------------------------------------------------------------------------- 1 | Paper: *Symbolic and Automatic Differentiation of Languages*. 2 | Appeared at ICFP 2021. 3 | 4 | [See latest version of the paper (PDF), video, and slides (PDF) here.](http://conal.net/papers/language-derivatives/) 5 | Comments appreciated as GitHub issues and pull requests or email (firstname@firstname.net). 6 | 7 | [Reviews](reviews.md) and [my responses](response-to-reviews.md) are included in the repository, along with the [submitted version](icfp21-submitted.pdf) to which the reviews & responses refer. 8 | 9 | Run "make" to compile the sources and typeset the paper. 10 | The Agda code should work at least with [Agda](https://github.com/agda/agda) 2.6.1.3 and [agda-stdlib](https://github.com/agda/agda-stdlib) version 1.5. 11 | 12 | **Abstract:** 13 | 14 | > Formal languages are usually defined in terms of set theory. Choosing type theory instead gives us languages as type-level predicates over strings. Applying a language to a string yields a type whose elements are language membership proofs describing *how* a string parses in the language. The usual building blocks of languages (including union, concatenation, and Kleene closure) have precise and compelling specifications uncomplicated by operational strategies and are easily generalized to a few general domain-transforming and codomain-transforming operations on predicates. 15 | > 16 | > A simple characterization of languages (and indeed functions from lists to any type) captures the essential idea behind language "differentiation" as used for recognizing languages, leading to a collection of lemmas about type-level predicates. 17 | > These lemmas are the heart of two dual parsing implementations---using (inductive) regular expressions and (coinductive) tries---each containing the same code but in dual arrangements (with representation and primitive operations trading places). 18 | > The regular expression version corresponds to symbolic differentiation, while the trie version corresponds to automatic differentiation. 19 | > 20 | > The relatively easy-to-prove properties of type-level languages transfer almost effortlessly to the decidable implementations. In particular, despite the inductive and coinductive nature of regular expressions and tries respectively, we need neither inductive nor coinductive/bisimulation arguments to prove algebraic properties. 21 | 22 | -------------------------------------------------------------------------------- /Language.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | {-# OPTIONS --safe --without-K #-} 3 | 4 | module Language {ℓ} (A : Set ℓ) where 5 | 6 | open import Level 7 | 8 | open import Algebra.Core 9 | -- open import Data.Empty 10 | -- open import Data.Unit 11 | open import Data.Sum 12 | open import Data.Product 13 | open import Data.List 14 | open import Data.List.Relation.Unary.All 15 | open import Relation.Binary.PropositionalEquality hiding ([_]) 16 | 17 | open import Misc {ℓ} 18 | 19 | Lang : Set (suc ℓ) 20 | ∅ : Lang 21 | 𝒰 : Lang 22 | _∪_ : Op₂ Lang 23 | _∩_ : Op₂ Lang 24 | _·_ : Set ℓ → Op₁ Lang 25 | 𝟏 : Lang 26 | _⋆_ : Op₂ Lang 27 | _☆ : Op₁ Lang 28 | ` : A → Lang 29 | 30 | infixr 6 _∪_ 31 | infixr 6 _∩_ 32 | infixl 7 _·_ 33 | infixl 7 _⋆_ 34 | infixl 10 _☆ 35 | \end{code} 36 | 37 | %<*Lang-ops> 38 | { 39 | \begin{center} 40 | \mathindent-20ex 41 | \begin{code} 42 | Lang = A ✶ → Set ℓ 43 | \end{code} 44 | \end{center} 45 | \ifacm\else 46 | \vspace{-3ex} 47 | \fi 48 | \iftalk \pause \fi 49 | \mathindent0ex 50 | \hfill 51 | \begin{minipage}[t]{20ex} 52 | \begin{code} 53 | ∅ w = ⊥ 54 | 55 | 𝒰 w = ⊤ 56 | 57 | 𝟏 w = w ≡ [] 58 | 59 | ` c w = w ≡ [ c ] 60 | 61 | (s · P) w = s × P w 62 | \end{code} 63 | \end{minipage} 64 | \hfill 65 | \begin{minipage}[t]{48ex} 66 | \begin{code} 67 | (P ∪ Q) w = P w ⊎ Q w 68 | 69 | (P ∩ Q) w = P w × Q w 70 | 71 | (P ⋆ Q) w = ∃⇃ λ (u , v) → (w ≡ u ⊙ v) × P u × Q v 72 | 73 | (P ☆) w = ∃ λ ws → (w ≡ concat ws) × All P ws 74 | \end{code} 75 | \end{minipage} 76 | \hfill\; 77 | } 78 | % 79 | 80 | \begin{code} 81 | module AltStar where 82 | infixl 10 _✪ 83 | data _✪ (P : Lang) : Lang where 84 | zero✪ : (P ✪) [] 85 | suc✪ : ∀ {w} → (P ⋆ P ✪) w → (P ✪) w 86 | \end{code} 87 | 88 | \begin{code} 89 | private 90 | module Stuff where 91 | private variable B X : Set ℓ 92 | \end{code} 93 | 94 | %<*foldr-concat> 95 | \begin{minipage}{18em} 96 | \begin{code} 97 | concat⇂ : (A ✶) ✶ → A ✶ 98 | concat⇂ = foldr _⊙_ [] 99 | \end{code} 100 | \end{minipage} 101 | %% where\hspace{3em} 102 | \begin{minipage}{18em} 103 | \begin{code} 104 | foldr⇂ : (B → X → X) → X → B ✶ → X 105 | foldr⇂ h x [] = x 106 | foldr⇂ h x (b ∷ bs) = h b (foldr⇂ h x bs) 107 | \end{code} 108 | \end{minipage} 109 | % 110 | 111 | %<*All> 112 | \begin{code}[hide] 113 | infixr 5 _∷_ 114 | \end{code} 115 | \begin{code} 116 | data All⇃ (P : B → Set ℓ) : B ✶ → Set ℓ where 117 | [] : All⇃ P [] 118 | _∷_ : ∀ {b bs} → P b → All P bs → All⇃ P (b ∷ bs) 119 | \end{code} 120 | % 121 | -------------------------------------------------------------------------------- /Closed/Instances.agda: -------------------------------------------------------------------------------- 1 | -- Some examples of closed semirings 2 | 3 | {-# OPTIONS --safe #-} 4 | {-# OPTIONS --without-K #-} 5 | 6 | module Closed.Instances where 7 | 8 | open import Level 9 | open import Algebra.Core 10 | open import Algebra.Structures 11 | open import Relation.Binary.PropositionalEquality 12 | import Relation.Binary.Reasoning.Setoid as SetoidR 13 | open import Function.Properties.Inverse using (↔-isEquivalence) 14 | 15 | open import Closed.Structures 16 | open import Closed.Bundles 17 | 18 | -- Booleans 19 | module _ where 20 | 21 | open import Data.Bool 22 | open import Data.Bool.Properties 23 | 24 | infixl 10 _✦ 25 | -- b ✦ ≡ true ∨ b ∧ b ✦ 26 | _✦ : Op₁ Bool 27 | _ ✦ = true 28 | 29 | ∨-∧-isClosedCommutativeSemiring : 30 | IsClosedCommutativeSemiring _≡_ _∨_ _∧_ false true _✦ 31 | ∨-∧-isClosedCommutativeSemiring = record 32 | { isCommutativeSemiring = ∨-∧-isCommutativeSemiring 33 | ; star = λ _ → refl 34 | } 35 | 36 | ∨-∧-closedCommutativeSemiring : ClosedCommutativeSemiring zero zero 37 | ∨-∧-closedCommutativeSemiring = 38 | record { isClosedCommutativeSemiring = ∨-∧-isClosedCommutativeSemiring } 39 | 40 | -- Types 41 | module Types {ℓ : Level} where 42 | 43 | open import Function.Bundles 44 | open import Data.Product 45 | open import Data.Sum 46 | open import Data.Unit.Polymorphic 47 | open import Data.Empty.Polymorphic 48 | open import Data.Product.Algebra 49 | open import Data.List renaming (List to _✶) 50 | 51 | -- A ✶ ↔ ⊤ ⊎ A × A ✶ 52 | ✶-star : Star {ℓ = ℓ} _↔_ _⊎_ _×_ ⊥ ⊤ _✶ 53 | ✶-star _ = mk↔′ 54 | (λ {[] → inj₁ tt ; (x ∷ xs) → inj₂ (x , xs)}) 55 | (λ {(inj₁ _) → [] ; (inj₂ (x , xs)) → x ∷ xs}) 56 | (λ {(inj₁ _) → refl ; (inj₂ (x , xs)) → refl}) 57 | (λ {[] → refl ; (x ∷ xs) → refl}) 58 | 59 | ⊎-×-isClosedCommutativeSemiring : 60 | IsClosedCommutativeSemiring {ℓ = ℓ} _↔_ _⊎_ _×_ ⊥ ⊤ _✶ 61 | ⊎-×-isClosedCommutativeSemiring = record 62 | { isCommutativeSemiring = ⊎-×-isCommutativeSemiring ℓ 63 | ; star = ✶-star 64 | } 65 | 66 | ×-⊎-closedCommutativeSemiring : ClosedCommutativeSemiring (suc ℓ) ℓ 67 | ×-⊎-closedCommutativeSemiring = 68 | record { isClosedCommutativeSemiring = ⊎-×-isClosedCommutativeSemiring } 69 | 70 | -- A ✶ is satisfiable for all A. If & when we extract more than one satisfying 71 | -- value, we can find more here. 72 | open import Relation.Nullary 73 | 74 | _✶-reflects : ∀ {A b} → Reflects {p = ℓ} A b → Reflects (A ✶) (b ✦) 75 | _ ✶-reflects = ofʸ [] 76 | 77 | infixl 10 _✶-dec 78 | _✶-dec : {A : Set ℓ} → Dec A → Dec (A ✶) 79 | does (a? ✶-dec) = does a? ✦ -- i.e., true 80 | proof (a? ✶-dec) = proof a? ✶-reflects 81 | 82 | -- TODO: relations using Relation.Binary.Construct.Closure.ReflexiveTransitive 83 | -------------------------------------------------------------------------------- /Symbolic.lagda: -------------------------------------------------------------------------------- 1 | \begin{code}[hide] 2 | 3 | open import Relation.Binary.PropositionalEquality using (_≡_) ; open _≡_ 4 | open import Decidability hiding (_◂_) 5 | 6 | module Symbolic {ℓ} {A : Set ℓ} (_≟_ : Decidable₂ {A = A} _≡_) where 7 | 8 | open import Level 9 | open import Data.List using ([]; _∷_) 10 | 11 | open import Misc {ℓ} 12 | open import Inverses {ℓ} 13 | 14 | module ◇ where 15 | -- open import Language A public 16 | open import Predicate public ; open ListOps A public 17 | open import Calculus A public 18 | 19 | open ◇ using (ν⋆; δ⋆; ν☆; δ☆; ν𝟏; δ𝟏; ν`; δ`) 20 | 21 | private 22 | variable 23 | P Q : ◇.Lang 24 | s : Set ℓ 25 | \end{code} 26 | 27 | %<*api> 28 | {\mathindent0ex 29 | \begin{center} 30 | \begin{code}[hide] 31 | infixr 6 _∪_ 32 | infixr 7 _∩_ 33 | infixl 7 _⋆_ 34 | infixr 7 _·_ 35 | infix 9 _◂_ 36 | infixl 10 _☆ 37 | \end{code} 38 | \begin{code} 39 | data Lang : ◇.Lang → Set (suc ℓ) where 40 | ∅ : Lang ◇.∅ 41 | 𝒰 : Lang ◇.𝒰 42 | _∪_ : Lang P → Lang Q → Lang (P ◇.∪ Q) 43 | _∩_ : Lang P → Lang Q → Lang (P ◇.∩ Q) 44 | _·_ : Dec s → Lang P → Lang (s ◇.· P) 45 | 𝟏 : Lang ◇.𝟏 46 | _⋆_ : Lang P → Lang Q → Lang (P ◇.⋆ Q) 47 | _☆ : Lang P → Lang (P ◇.☆) 48 | ` : (a : A) → Lang (◇.` a) 49 | _◂_ : (Q ⟷ P) → Lang P → Lang Q 50 | \end{code} 51 | \end{center} 52 | \iftalk 53 | \vspace{-3.5ex} 54 | \fi 55 | \hfill 56 | \begin{minipage}[c]{33ex} 57 | \begin{code} 58 | ν : Lang P → Dec (◇.ν P) 59 | δ : Lang P → (a : A) → Lang (◇.δ P a) 60 | \end{code} 61 | \end{minipage} 62 | \hfill 63 | \begin{minipage}[c]{25ex} 64 | \begin{code} 65 | ⟦_⟧‽ : Lang P → Decidable P 66 | ⟦ p ⟧‽ [] = ν p 67 | ⟦ p ⟧‽ (a ∷ w) = ⟦ δ p a ⟧‽ w 68 | \end{code} 69 | \end{minipage} 70 | \hfill\; 71 | } 72 | % 73 | 74 | %<*semantics> 75 | \begin{code} 76 | ⟦_⟧ : Lang P → ◇.Lang 77 | ⟦_⟧ {P} r = P 78 | \end{code} 79 | % 80 | 81 | %<*defs> 82 | {\mathindent0ex 83 | \begin{code}[hide] 84 | private 85 | variable 86 | p : Lang P 87 | q : Lang Q 88 | 89 | \end{code} 90 | \setstretch{1.6} 91 | \hfill 92 | %\hspace{-1.2ex}%% To align with Automatic. Why different? 93 | \begin{minipage}{30ex} 94 | \begin{code} 95 | ν ∅ = ⊥‽ 96 | ν 𝒰 = ⊤‽ 97 | ν (p ∪ q) = ν p ⊎‽ ν q 98 | ν (p ∩ q) = ν p ×‽ ν q 99 | ν (s · p) = s ×‽ ν p 100 | ν 𝟏 = ν𝟏 ◃ ⊤‽ 101 | ν (p ⋆ q) = ν⋆ ◃ (ν p ×‽ ν q) 102 | ν (p ☆) = ν☆ ◃ (ν p ✶‽) 103 | ν (` a) = ν` ◃ ⊥‽ 104 | ν (f ◂ p) = f ◃ ν p 105 | \end{code} 106 | \end{minipage} 107 | \hfill 108 | \begin{minipage}{38ex} 109 | \begin{code} 110 | δ ∅ a = ∅ 111 | δ 𝒰 a = 𝒰 112 | δ (p ∪ q) a = δ p a ∪ δ q a 113 | δ (p ∩ q) a = δ p a ∩ δ q a 114 | δ (s · p) a = s · δ p a 115 | δ 𝟏 a = δ𝟏 ◂ ∅ 116 | δ (p ⋆ q) a = δ⋆ ◂ (ν p · δ q a ∪ δ p a ⋆ q) 117 | δ (p ☆) a = δ☆ ◂ (ν p ✶‽ · (δ p a ⋆ p ☆)) 118 | δ (` c) a = δ` ◂ ((a ≟ c) · 𝟏) 119 | δ (f ◂ p) a = f ◂ δ p a 120 | \end{code} 121 | \end{minipage} 122 | \hfill\; 123 | } 124 | % 125 | -------------------------------------------------------------------------------- /Transport.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | {-# OPTIONS --guardedness --sized-types #-} 3 | 4 | open import Relation.Binary.PropositionalEquality using (_≡_) ; open _≡_ 5 | open import Decidability 6 | 7 | module Transport {ℓ} {A : Set ℓ} (_≟_ : Decidable₂ {A = A} _≡_) where 8 | 9 | open import Level 10 | open import Algebra 11 | open import Data.Product 12 | 13 | open import Existential ; open Inj 14 | open import Closed 15 | 16 | module ◇ where 17 | open import Predicate public 18 | open ListOps A public 19 | 20 | module □ where 21 | open import Symbolic _≟_ public 22 | 23 | \end{code} 24 | 25 | \begin{code} 26 | module Examples where 27 | 28 | open import Data.List.Properties 29 | 30 | open import Predicate.Algebra (++-isMonoid {A = A}) 31 | 32 | open import Size 33 | 34 | open import Closed.Instances ; open Types {ℓ} 35 | 36 | module Symbolic where 37 | open import Symbolic _≟_ public 38 | 39 | module Automatic where 40 | open import SizedAutomatic _≟_ public 41 | 42 | decCCS : ClosedCommutativeSemiring (suc ℓ) ℓ 43 | symbolicCS₁ : CommutativeSemiring (suc ℓ) ℓ 44 | symbolicCS₂ : ClosedSemiring (suc ℓ) ℓ 45 | automaticCS₁ : CommutativeSemiring (suc ℓ) ℓ 46 | automaticCS₂ : ClosedSemiring (suc ℓ) ℓ 47 | \end{code} 48 | %<*examples> 49 | \mathindent0ex 50 | \begin{code} 51 | decCCS = mkClosedCommutativeSemiring ×-⊎-closedCommutativeSemiring 52 | Dec _⊎‽_ _×‽_ ⊥‽ ⊤‽ _✶‽ 53 | 54 | symbolicCS₁ = mkCommutativeSemiring (∩-∪-commutativeSemiring _) Lang _∪_ _∩_ ∅ 𝒰 55 | where open Symbolic 56 | 57 | symbolicCS₂ = mkClosedSemiring ⋆-∪-closedSemiring Lang _∪_ _⋆_ ∅ 𝟏 _☆ 58 | where open Symbolic 59 | 60 | automaticCS₁ = mkCommutativeSemiring (∩-∪-commutativeSemiring _) (Lang ∞) _∪_ _∩_ ∅ 𝒰 61 | where open Automatic 62 | 63 | automaticCS₂ = mkClosedSemiring ⋆-∪-closedSemiring (Lang ∞) _∪_ _⋆_ ∅ 𝟏 _☆ 64 | where open Automatic 65 | \end{code} 66 | % 67 | 68 | \begin{code} 69 | module Wrap where 70 | Lang : Set (suc ℓ) 71 | 72 | infixr 6 _∪_ 73 | infixr 6 _∩_ 74 | infixl 7 _·_ 75 | infixl 7 _⋆_ 76 | infixl 10 _☆ 77 | 78 | ∅ : Lang 79 | 𝒰 : Lang 80 | _∪_ : Op₂ Lang 81 | _∩_ : Op₂ Lang 82 | 𝟏 : Lang 83 | _⋆_ : Op₂ Lang 84 | _·_ : ∀ {s} → Dec s → Op₁ Lang 85 | ` : A → Lang 86 | _☆ : Op₁ Lang 87 | 88 | Lang = ∃ □.Lang 89 | 90 | ∅ = inj □.∅ 91 | 𝒰 = inj □.𝒰 92 | _∪_ = inj₂ □._∪_ 93 | _∩_ = inj₂ □._∩_ 94 | 𝟏 = inj □.𝟏 95 | _⋆_ = inj₂ □._⋆_ 96 | _·_ s = inj₁ (s □.·_) 97 | ` c = inj (□.` c) 98 | _☆ = inj₁ □._☆ 99 | \end{code} 100 | 101 | \begin{code} 102 | module ManualWrap where 103 | \end{code} 104 | %<*wrapped-Lang> 105 | \begin{code} 106 | Lang = ∃ □.Lang 107 | \end{code} 108 | % 109 | %<*wrapped-union> 110 | \begin{code} 111 | _∪_ : Lang → Lang → Lang 112 | (P , p) ∪ (Q , q) = P ◇.∪ Q , p □.∪ q 113 | \end{code} 114 | % (_ , p) ∪ (_ , q) = (_ , p □.∪ q) 115 | 116 | % 117 | -------------------------------------------------------------------------------- /macros.tex: -------------------------------------------------------------------------------- 1 | \RequirePackage{xcolor} 2 | 3 | \newcommand\nc\newcommand 4 | \nc\rnc\renewcommand 5 | 6 | \nc\out[1]{} 7 | 8 | %% \nc\noteOut[2]{\note{#1}\out{#2}} 9 | 10 | %% To redefine for a non-draft 11 | \nc\indraft[1]{#1} 12 | 13 | %% I think \note gets defined by beamer. 14 | \let\note\undefined 15 | 16 | \nc\note[1]{\indraft{\textcolor{red}{#1}}} 17 | 18 | \nc\notefoot[1]{\note{\footnote{\note{#1}}}} 19 | 20 | \nc\todo[1]{\note{To do: #1}} 21 | 22 | \nc\eqnlabel[1]{\label{equation:#1}} 23 | \nc\eqnref[1]{Equation~\ref{equation:#1}} 24 | \nc\eqnreftwo[2]{Equations~\ref{equation:#1} and \ref{equation:#2}} 25 | 26 | \nc\figlabel[1]{\label{fig:#1}} 27 | \nc\figref[1]{Figure~\ref{fig:#1}} 28 | \nc\figreftwo[2]{Figures~\ref{fig:#1} and \ref{fig:#2}} 29 | 30 | \nc\seclabel[1]{\label{sec:#1}} 31 | \nc\secref[1]{Section~\ref{sec:#1}} 32 | \nc\secreftwo[2]{Sections~\ref{sec:#1} and~\ref{sec:#2}} 33 | \nc\secrefs[2]{Sections \ref{sec:#1} through \ref{sec:#2}} 34 | 35 | \nc\appref[1]{Appendix~\ref{sec:#1}} 36 | 37 | %% The name \secdef is already taken 38 | \nc\sectiondef[1]{\section{#1}\seclabel{#1}} 39 | \nc\subsectiondef[1]{\subsection{#1}\seclabel{#1}} 40 | \nc\subsubsectiondef[1]{\subsubsection{#1}\seclabel{#1}} 41 | 42 | \nc\needcite{\note{[ref]}} 43 | 44 | \nc\sectionl[1]{\section{#1}\seclabel{#1}} 45 | \nc\subsectionl[1]{\subsection{#1}\seclabel{#1}} 46 | 47 | \nc\workingHere{ 48 | \vspace{1ex} 49 | \begin{center} 50 | \setlength{\fboxsep}{3ex} 51 | \setlength{\fboxrule}{4pt} 52 | \huge\textcolor{red}{\framebox{Working here}} 53 | \end{center} 54 | \vspace{1ex} 55 | } 56 | 57 | %% For multiple footnotes at a point. Adapted to recognize \notefoot as well 58 | %% as \footnote. See https://tex.stackexchange.com/a/71347, 59 | \let\oldFootnote\footnote 60 | \nc\nextToken\relax 61 | \rnc\footnote[1]{% 62 | \oldFootnote{#1}\futurelet\nextToken\isFootnote} 63 | \nc\footcomma[1]{\ifx#1\nextToken\textsuperscript{,}\fi} 64 | \nc\isFootnote{% 65 | \footcomma\footnote 66 | \footcomma\notefoot 67 | } 68 | 69 | % Arguments: env, label, caption, body 70 | \nc\figdefG[4]{\begin{#1}%[tbp] 71 | \begin{center} 72 | #4 73 | \end{center} 74 | \vspace{-1ex} 75 | \caption{#3} 76 | \figlabel{#2} 77 | \end{#1}} 78 | 79 | % Arguments: label, caption, body 80 | \nc\figdef{\figdefG{figure}} 81 | \nc\figdefwide{\figdefG{figure*}} 82 | 83 | % Arguments: label, caption, body 84 | \nc\figrefdef[3]{\figref{#1}\figdef{#1}{#2}{#3}} 85 | 86 | \usepackage{hyperref} 87 | 88 | \nc\stdlibCitet[1]{\citet[\AM{#1}]{agda-stdlib}} 89 | \nc\stdlibCitep[1]{\citep[\AM{#1}]{agda-stdlib}} 90 | \nc\stdlibCite\stdlibCitep 91 | 92 | %% https://tex.stackexchange.com/questions/94699/absolutely-definitely-preventing-page-break 93 | 94 | \renewenvironment{samepage} 95 | {\par\nobreak\vfil\penalty0\vfilneg 96 | \vtop\bgroup} 97 | {\par\xdef\tpd{\the\prevdepth}\egroup 98 | \prevdepth=\tpd} 99 | 100 | \nc\source{None} 101 | 102 | \nc\agda[1]{\ExecuteMetaData[\source.tex]{#1}} 103 | 104 | %% \nc\agda[1]{\begin{samepage}\ExecuteMetaData[\source.tex]{#1}\end{samepage}} 105 | -------------------------------------------------------------------------------- /Decidability.lagda: -------------------------------------------------------------------------------- 1 | %% \section{Decidability} 2 | 3 | \begin{code}[hide] 4 | {-# OPTIONS --safe #-} 5 | 6 | open import Level 7 | 8 | module Decidability {ℓ : Level} where 9 | 10 | open import Function using (_∘_; _↔_; Inverse) 11 | open import Data.List using ([]) 12 | 13 | open import Misc {ℓ} 14 | open import Inverses {ℓ} 15 | 16 | private 17 | variable 18 | A B X : Set ℓ 19 | P Q : A → Set ℓ 20 | \end{code} 21 | 22 | %<*Dec> 23 | \AgdaTarget{Dec; yes; no} 24 | \begin{code} 25 | data Dec (X : Set ℓ) : Set ℓ where 26 | yes : X → Dec X 27 | no : ¬ X → Dec X 28 | \end{code}% 29 | % 30 | 31 | %<*¬> 32 | \begin{code}[hide] 33 | private 34 | module Stuff where 35 | ¬⇃_ : Set ℓ → Set ℓ 36 | \end{code} 37 | \begin{code}[inline] 38 | ¬⇃ X = X → ⊥ 39 | \end{code} 40 | %¬> 41 | 42 | %<*Decidable> 43 | \AgdaTarget{Decidable} 44 | \begin{code}[hide] 45 | Decidable : (X → Set ℓ) → Set ℓ 46 | \end{code} 47 | \begin{code} 48 | Decidable P = ∀ x → Dec (P x) 49 | \end{code} 50 | % 51 | 52 | \AgdaTarget{Decidable₂} 53 | \begin{code} 54 | Decidable₂ : (A → B → Set ℓ) → Set ℓ 55 | Decidable₂ _∼_ = ∀ a b → Dec (a ∼ b) 56 | \end{code} 57 | 58 | %<*isomorphisms> 59 | \begin{code}[hide] 60 | open import Function.Equality using (Π) ; open Π 61 | open import Function.Equivalence using (Equivalence; equivalence; _⇔_) 62 | 63 | map′ : (A → B) → (B → A) → Dec A → Dec B 64 | map′ A→B B→A (yes a) = yes (A→B a) 65 | map′ A→B B→A (no ¬a) = no (¬a ∘ B→A) 66 | 67 | map‽⇔ : A ⇔ B → Dec A → Dec B 68 | map‽⇔ A⇔B = map′ (to ⟨$⟩_) (from ⟨$⟩_) where open Equivalence A⇔B 69 | \end{code} 70 | \AgdaTarget{\_▹\_, \_◃\_} 71 | \begin{code}[hide] 72 | infixr 9 _▹_ 73 | infixr 9 _◃_ 74 | infixr 9 _▸_ 75 | infixr 9 _◂_ 76 | 77 | _▹_ : A ↔ B → Dec A → Dec B 78 | f ▹ a? = map‽⇔ (↔→⇔ f) a? 79 | 80 | _▸_ : (P ⟷ Q) → Decidable P → Decidable Q 81 | (f ▸ p?) w = f ▹ p? w 82 | \end{code} 83 | \begin{code} 84 | _◃_ : B ↔ A → Dec A → Dec B 85 | 86 | _◂_ : Q ⟷ P → Decidable P → Decidable Q 87 | \end{code} 88 | \begin{code}[hide] 89 | g ◃ a? = ↔Eq.sym g ▹ a? 90 | 91 | (g ◂ p?) w = g ◃ p? w 92 | \end{code} 93 | % 94 | 95 | %<*compositional-dec> 96 | {\mathindent0ex 97 | \begin{code}[hide] 98 | open import Data.Sum 99 | open import Data.Product 100 | open import Function using (id; _∘_) 101 | 102 | infixr 1 _⊎‽_ 103 | infixr 2 _×‽_ 104 | \end{code} 105 | \hfill\; 106 | \begin{minipage}{37ex} 107 | \begin{code} 108 | ⊥‽ : Dec ⊥ 109 | ⊥‽ = no (λ ()) 110 | 111 | 112 | _⊎‽_ : Dec A → Dec B → Dec (A ⊎ B) 113 | no ¬a ⊎‽ no ¬b = no [ ¬a , ¬b ] 114 | yes a ⊎‽ no ¬b = yes (inj₁ a) 115 | _ ⊎‽ yes b = yes (inj₂ b) 116 | \end{code} 117 | \end{minipage} 118 | \hfill 119 | \begin{minipage}{37ex} 120 | \begin{code} 121 | ⊤‽ : Dec ⊤ 122 | ⊤‽ = yes tt 123 | 124 | 125 | _×‽_ : Dec A → Dec B → Dec (A × B) 126 | yes a ×‽ yes b = yes (a , b) 127 | no ¬a ×‽ yes b = no (¬a ∘ proj₁) 128 | _ ×‽ no ¬b = no (¬b ∘ proj₂) 129 | \end{code} 130 | \end{minipage} 131 | \hfill\; 132 | \begin{code}[hide] 133 | ¬‽_ : Dec A → Dec (¬ A) 134 | ¬‽ yes a = no λ ¬a → ¬a a 135 | ¬‽ no ¬a = yes λ a → ¬a a 136 | \end{code} 137 | \begin{center} 138 | \begin{code} 139 | _✶‽ : Dec A → Dec (A ✶) 140 | _ ✶‽ = yes [] 141 | \end{code} 142 | \end{center} 143 | } 144 | % 145 | -------------------------------------------------------------------------------- /Predicate/Algebra.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --safe #-} 2 | -- {-# OPTIONS --without-K #-} 3 | 4 | open import Level 5 | open import Algebra.Core 6 | open import Relation.Binary.PropositionalEquality 7 | using (_≡_; _≗_; cong; subst) renaming (refl to refl≡; sym to sym≡) 8 | import Algebra.Structures as S 9 | 10 | module Predicate.Algebra {ℓ : Level} {M : Set ℓ} {_∙_ : Op₂ M} {ε : M} 11 | (isMonoid-M : S.IsMonoid _≡_ _∙_ ε) where 12 | 13 | private 14 | variable 15 | A B C : Set ℓ 16 | 17 | open import Function using (mk↔′) 18 | open import Algebra.Definitions 19 | open import Algebra.Bundles 20 | open import Algebra.Module.Bundles 21 | open import Data.Product.Algebra using (×-⊎-commutativeSemiring) 22 | 23 | -- Local 24 | import Algebra.Construct.Function as C 25 | open import Inverses {ℓ} 26 | open import Predicate {ℓ} 27 | 28 | open S.IsMonoid isMonoid-M 29 | 30 | open import Algebra.Structures {suc ℓ} {ℓ} {Pred M} _⟷_ 31 | 32 | open MonoidOps _∙_ ε 33 | 34 | -- Pred as commutative semiring 35 | ∩-∪-commutativeSemiring : Set ℓ → CommutativeSemiring _ _ 36 | ∩-∪-commutativeSemiring A = C.commutativeSemiring A (×-⊎-commutativeSemiring ℓ) 37 | 38 | open import Predicate.Properties {ℓ} 39 | open Pred-Semimodule 40 | 41 | open MonoidSemiringProperties isMonoid-M 42 | 43 | ⋆-isMagma : IsMagma _⋆_ 44 | ⋆-isMagma = record { isEquivalence = ⟷-isEquivalence ; ∙-cong = ⋆-cong } 45 | 46 | ⋆-isSemigroup : IsSemigroup _⋆_ 47 | ⋆-isSemigroup = record { isMagma = ⋆-isMagma ; assoc = mapⱽ-assoc assoc } 48 | 49 | ⋆-isMonoid : IsMonoid _⋆_ 𝟏 50 | ⋆-isMonoid = record 51 | { isSemigroup = ⋆-isSemigroup 52 | ; identity = mapⱽ-identity identity 53 | } 54 | 55 | ⋆-∪-isSemiringWithoutAnnihilatingZero : 56 | IsSemiringWithoutAnnihilatingZero _∪_ _⋆_ ∅ 𝟏 57 | ⋆-∪-isSemiringWithoutAnnihilatingZero = record 58 | { +-isCommutativeMonoid = +ᴹ-isCommutativeMonoid {M} 59 | ; *-isMonoid = ⋆-isMonoid 60 | ; distrib = mapⱽ-distrib 61 | } 62 | 63 | ⋆-∪-isSemiring : IsSemiring _∪_ _⋆_ ∅ 𝟏 64 | ⋆-∪-isSemiring = record 65 | { isSemiringWithoutAnnihilatingZero = ⋆-∪-isSemiringWithoutAnnihilatingZero 66 | ; zero = mapⱽ-zero 67 | } 68 | 69 | ⋆-∪-isCommutativeSemiring : 70 | Commutative _≡_ _∙_ → IsCommutativeSemiring _∪_ _⋆_ ∅ 𝟏 71 | ⋆-∪-isCommutativeSemiring ∙-comm = record 72 | { isSemiring = ⋆-∪-isSemiring 73 | ; *-comm = mapⱽ-comm ∙-comm 74 | } 75 | 76 | -- TODO: consider parametrizing each algebraic structure by the corresponding 77 | -- algebraic structure on M. 78 | 79 | open import Closed ; open Closed.Types {ℓ} 80 | 81 | ⋆-∪-isClosedSemiring : IsClosedSemiring _⟷_ _∪_ _⋆_ ∅ 𝟏 _☆ 82 | ⋆-∪-isClosedSemiring = record { isSemiring = ⋆-∪-isSemiring 83 | ; star = λ _ → ☆-star 84 | } 85 | 86 | ⋆-∪-isClosedCommutativeSemiring : 87 | Commutative _≡_ _∙_ → IsClosedCommutativeSemiring _⟷_ _∪_ _⋆_ ∅ 𝟏 _☆ 88 | ⋆-∪-isClosedCommutativeSemiring ∙-comm = record 89 | { isCommutativeSemiring = ⋆-∪-isCommutativeSemiring ∙-comm 90 | ; star = λ _ → ☆-star 91 | } 92 | 93 | ⋆-∪-closedSemiring : ClosedSemiring _ _ 94 | ⋆-∪-closedSemiring = record { isClosedSemiring = ⋆-∪-isClosedSemiring } 95 | 96 | ⋆-∪-closedCommutativeSemiring : Commutative _≡_ _∙_ → ClosedCommutativeSemiring _ _ 97 | ⋆-∪-closedCommutativeSemiring ∙-comm = 98 | record { isClosedCommutativeSemiring = ⋆-∪-isClosedCommutativeSemiring ∙-comm } 99 | -------------------------------------------------------------------------------- /SizedAutomatic.lagda: -------------------------------------------------------------------------------- 1 | \begin{code}[hide] 2 | {-# OPTIONS --guardedness --sized-types #-} 3 | 4 | open import Level 5 | open import Decidability hiding (_◂_) 6 | open import Relation.Binary.PropositionalEquality using (_≡_) ; open _≡_ 7 | 8 | module SizedAutomatic {ℓ} {A : Set ℓ} (_≟_ : Decidable₂ {A = A} _≡_) where 9 | 10 | open import Size 11 | open import Data.List using ([]; _∷_) 12 | 13 | open import Misc {ℓ} 14 | open import Inverses {ℓ} 15 | 16 | module ◇ where 17 | -- open import Language A public 18 | open import Predicate public ; open ListOps A public 19 | open import Calculus A public 20 | 21 | open ◇ using (ν⋆; δ⋆; ν☆; δ☆; ν𝟏; δ𝟏; ν`; δ`) 22 | 23 | private 24 | variable 25 | P Q : ◇.Lang 26 | s : Set ℓ 27 | i : Size 28 | \end{code} 29 | 30 | %<*api> 31 | {\mathindent0ex 32 | \hfill 33 | \begin{minipage}[c]{52ex} 34 | \begin{code}[hide] 35 | infixr 6 _∪_ 36 | infixl 7 _∩_ 37 | infixl 7 _⋆_ 38 | infixr 7 _·_ 39 | infix 9 _◂_ 40 | infixl 10 _☆ 41 | \end{code} 42 | \begin{code} 43 | record Lang i (P : ◇.Lang) : Set (suc ℓ) where 44 | coinductive 45 | field 46 | ν : Dec (◇.ν P) 47 | δ : ∀ {j : Size< i} → (a : A) → Lang j (◇.δ P a) 48 | \end{code} 49 | \begin{code}[hide] 50 | open Lang 51 | \end{code} 52 | \end{minipage} 53 | \hfill 54 | \begin{minipage}[c]{27ex} 55 | \begin{code} 56 | ⟦_⟧‽ : Lang ∞ P → Decidable P 57 | ⟦ p ⟧‽ [] = ν p 58 | ⟦ p ⟧‽ (a ∷ w) = ⟦ δ p a ⟧‽ w 59 | \end{code} 60 | \end{minipage} 61 | \hfill\; 62 | \iftalk 63 | \vspace{-3ex} 64 | \fi 65 | \begin{center} 66 | \begin{code} 67 | ∅ : Lang i ◇.∅ 68 | 𝒰 : Lang i ◇.𝒰 69 | _∪_ : Lang i P → Lang i Q → Lang i (P ◇.∪ Q) 70 | _∩_ : Lang i P → Lang i Q → Lang i (P ◇.∩ Q) 71 | _·_ : Dec s → Lang i P → Lang i (s ◇.· P) 72 | 𝟏 : Lang i ◇.𝟏 73 | _⋆_ : Lang i P → Lang i Q → Lang i (P ◇.⋆ Q) 74 | _☆ : Lang i P → Lang i (P ◇.☆) 75 | ` : (a : A) → Lang i (◇.` a) 76 | _◂_ : (Q ⟷ P) → Lang i P → Lang i Q 77 | \end{code} 78 | \end{center} 79 | } 80 | % 81 | 82 | %<*semantics> 83 | \begin{code} 84 | ⟦_⟧ : Lang ∞ P → ◇.Lang 85 | ⟦_⟧ {P} _ = P 86 | \end{code} 87 | % 88 | 89 | %<*defs> 90 | {\mathindent0ex 91 | \setstretch{1.6} 92 | \rules{\begin{code} 93 | ν ∅ = ⊥‽ 94 | \end{code} 95 | }{\begin{code} 96 | δ ∅ a = ∅ 97 | \end{code}} 98 | 99 | \rules{\begin{code} 100 | ν 𝒰 = ⊤‽ 101 | \end{code} 102 | }{\begin{code} 103 | δ 𝒰 a = 𝒰 104 | \end{code}} 105 | 106 | \rules{\begin{code} 107 | ν (p ∪ q) = ν p ⊎‽ ν q 108 | \end{code} 109 | }{\begin{code} 110 | δ (p ∪ q) a = δ p a ∪ δ q a 111 | \end{code}} 112 | 113 | \rules{\begin{code} 114 | ν (p ∩ q) = ν p ×‽ ν q 115 | \end{code} 116 | }{\begin{code} 117 | δ (p ∩ q) a = δ p a ∩ δ q a 118 | \end{code}} 119 | 120 | \rules{\begin{code} 121 | ν (s · p) = s ×‽ ν p 122 | \end{code} 123 | }{\begin{code} 124 | δ (s · p) a = s · δ p a 125 | \end{code}} 126 | 127 | \rules{\begin{code} 128 | ν 𝟏 = ν𝟏 ◃ ⊤‽ 129 | \end{code} 130 | }{\begin{code} 131 | δ 𝟏 a = δ𝟏 ◂ ∅ 132 | \end{code}} 133 | 134 | \rules{\begin{code} 135 | ν (p ⋆ q) = ν⋆ ◃ (ν p ×‽ ν q) 136 | \end{code} 137 | }{\begin{code} 138 | δ (p ⋆ q) a = δ⋆ ◂ (ν p · δ q a ∪ δ p a ⋆ q) 139 | \end{code}} 140 | 141 | \rules{\begin{code} 142 | ν (p ☆) = ν☆ ◃ (ν p ✶‽) 143 | \end{code} 144 | }{\begin{code} 145 | δ (p ☆) a = δ☆ ◂ (ν p ✶‽ · (δ p a ⋆ p ☆)) 146 | \end{code}} 147 | 148 | \rules{\begin{code} 149 | ν (` a) = ν` ◃ ⊥‽ 150 | \end{code} 151 | }{\begin{code} 152 | δ (` c) a = δ` ◂ ((a ≟ c) · 𝟏) 153 | \end{code}} 154 | 155 | \rules{\begin{code} 156 | ν (f ◂ p) = f ◃ ν p 157 | \end{code} 158 | }{\begin{code} 159 | δ (f ◂ p) a = f ◂ δ p a 160 | \end{code}} 161 | } 162 | % 163 | -------------------------------------------------------------------------------- /Automatic.lagda: -------------------------------------------------------------------------------- 1 | \begin{code}[hide] 2 | {-# OPTIONS --guardedness #-} 3 | open import Level 4 | open import Decidability hiding (_◂_) 5 | open import Relation.Binary.PropositionalEquality using (_≡_) ; open _≡_ 6 | 7 | module Automatic {ℓ} {A : Set ℓ} (_≟_ : Decidable₂ {A = A} _≡_) where 8 | 9 | open import Data.List using ([]; _∷_) 10 | 11 | open import Misc {ℓ} 12 | open import Inverses {ℓ} 13 | 14 | module ◇ where 15 | -- open import Language A public 16 | open import Predicate public ; open ListOps A public 17 | open import Calculus A public 18 | 19 | open ◇ using (ν⋆; δ⋆; ν☆; δ☆; ν𝟏; δ𝟏; ν`; δ`) 20 | 21 | private 22 | variable 23 | P Q : ◇.Lang 24 | s : Set ℓ 25 | \end{code} 26 | 27 | %<*api> 28 | {\mathindent0ex 29 | \hfill 30 | \begin{minipage}[c]{37ex} 31 | \begin{code}[hide] 32 | infixr 6 _∪_ 33 | infixl 7 _∩_ 34 | infixl 7 _⋆_ 35 | infixr 7 _·_ 36 | infix 9 _◂_ 37 | infixl 10 _☆ 38 | \end{code} 39 | \begin{code} 40 | record Lang (P : ◇.Lang) : Set (suc ℓ) where 41 | coinductive 42 | field 43 | ν : Dec (◇.ν P) 44 | δ : (a : A) → Lang (◇.δ P a) 45 | \end{code} 46 | \begin{code}[hide] 47 | open Lang 48 | \end{code} 49 | \end{minipage} 50 | \hfill 51 | \begin{minipage}[c]{25ex} 52 | \begin{code} 53 | ⟦_⟧‽ : Lang P → Decidable P 54 | ⟦ p ⟧‽ [] = ν p 55 | ⟦ p ⟧‽ (a ∷ w) = ⟦ δ p a ⟧‽ w 56 | \end{code} 57 | \end{minipage} 58 | \hfill\; 59 | \iftalk 60 | \vspace{-3ex} 61 | \fi 62 | \begin{center} 63 | \begin{code} 64 | ∅ : Lang ◇.∅ 65 | 𝒰 : Lang ◇.𝒰 66 | _∪_ : Lang P → Lang Q → Lang (P ◇.∪ Q) 67 | _∩_ : Lang P → Lang Q → Lang (P ◇.∩ Q) 68 | _·_ : Dec s → Lang P → Lang (s ◇.· P) 69 | 𝟏 : Lang ◇.𝟏 70 | _⋆_ : Lang P → Lang Q → Lang (P ◇.⋆ Q) 71 | _☆ : Lang P → Lang (P ◇.☆) 72 | ` : (a : A) → Lang (◇.` a) 73 | _◂_ : (Q ⟷ P) → Lang P → Lang Q 74 | \end{code} 75 | \end{center} 76 | } 77 | % 78 | 79 | %<*semantics> 80 | \begin{code} 81 | ⟦_⟧ : Lang P → ◇.Lang 82 | ⟦_⟧ {P} _ = P 83 | \end{code} 84 | % 85 | 86 | %<*rules> 87 | \newcommand\rulesSep{-4.57ex} 88 | \newcommand\rules[2]{ 89 | \vspace{\rulesSep} 90 | \hspace{0ex}\\% Weirdness fix 91 | \hspace{0.00ex} 92 | \hfill 93 | \begin{minipage}{30ex} 94 | #1 95 | \end{minipage} 96 | \hfill 97 | \begin{minipage}{38ex} 98 | #2 99 | \end{minipage} 100 | \hfill\; 101 | } 102 | % 103 | 104 | %<*defs> 105 | {\mathindent0ex 106 | \setstretch{1.6} 107 | \rules{\begin{code} 108 | ν ∅ = ⊥‽ 109 | \end{code} 110 | }{\begin{code} 111 | δ ∅ a = ∅ 112 | \end{code}} 113 | 114 | \rules{\begin{code} 115 | ν 𝒰 = ⊤‽ 116 | \end{code} 117 | }{\begin{code} 118 | δ 𝒰 a = 𝒰 119 | \end{code}} 120 | 121 | \rules{\begin{code} 122 | ν (p ∪ q) = ν p ⊎‽ ν q 123 | \end{code} 124 | }{\begin{code} 125 | δ (p ∪ q) a = δ p a ∪ δ q a 126 | \end{code}} 127 | 128 | \rules{\begin{code} 129 | ν (p ∩ q) = ν p ×‽ ν q 130 | \end{code} 131 | }{\begin{code} 132 | δ (p ∩ q) a = δ p a ∩ δ q a 133 | \end{code}} 134 | 135 | \rules{\begin{code} 136 | ν (s · p) = s ×‽ ν p 137 | \end{code} 138 | }{\begin{code} 139 | δ (s · p) a = s · δ p a 140 | \end{code}} 141 | 142 | \rules{\begin{code} 143 | ν 𝟏 = ν𝟏 ◃ ⊤‽ 144 | \end{code} 145 | }{\begin{code} 146 | δ 𝟏 a = δ𝟏 ◂ ∅ 147 | \end{code}} 148 | 149 | \begin{code}[hide] 150 | {-# TERMINATING #-} 151 | \end{code} 152 | \rules{\begin{code} 153 | ν (p ⋆ q) = ν⋆ ◃ (ν p ×‽ ν q) 154 | \end{code} 155 | }{\begin{code} 156 | δ (p ⋆ q) a = δ⋆ ◂ (ν p · δ q a ∪ δ p a ⋆ q) 157 | \end{code}} 158 | 159 | \rules{\begin{code} 160 | ν (p ☆) = ν☆ ◃ (ν p ✶‽) 161 | \end{code} 162 | }{\begin{code} 163 | δ (p ☆) a = δ☆ ◂ (ν p ✶‽ · (δ p a ⋆ p ☆)) 164 | \end{code}} 165 | 166 | \rules{\begin{code} 167 | ν (` a) = ν` ◃ ⊥‽ 168 | \end{code} 169 | }{\begin{code} 170 | δ (` c) a = δ` ◂ ((a ≟ c) · 𝟏) 171 | \end{code}} 172 | 173 | \rules{\begin{code} 174 | ν (f ◂ p) = f ◃ ν p 175 | \end{code} 176 | }{\begin{code} 177 | δ (f ◂ p) a = f ◂ δ p a 178 | \end{code}} 179 | } 180 | % 181 | 182 | %<*termination> 183 | \begin{code} 184 | {-# TERMINATING #-} -- ? 185 | _⋆⇃_ : Lang P → Lang Q → Lang (P ◇.⋆ Q) 186 | ν (p ⋆⇃ q) = ν⋆ ◃ (ν p ×‽ ν q) 187 | δ (p ⋆⇃ q) a = δ⋆ ◂ (ν p · δ q a ∪ δ p a ⋆⇃ q) 188 | \end{code} 189 | % 190 | -------------------------------------------------------------------------------- /final-review-responses.md: -------------------------------------------------------------------------------- 1 | # ICFP 2021 Paper #87 final review responses 2 | 3 | From the final review: 4 |
5 | The reviewers want to thank the author for this interesting submission and are happy to accept it with the understanding that the author will make all of the following changes: 6 | 7 | 1. Clarify what's new in the context of proof mechanization or type theory, what's new in the context of parsing, and what's new in the context of formal languages & automata. In particular, some contributions do not seem to be novel in the last context. The author did not respond to Reviewer C's comment about the work of Lombardy and Sakarovitch, but their work seems to be similar to the derivative operations in the paper and the author should pay attention to it (in addition to other related works). 8 | 9 | 2. Explain the Agda syntax more for the readers who are not familiar with Agda. For example, add concrete examples showing how operators in Fig. 1 apply to concrete regular expressions and words, including the fully expanded types and values (as witnesses of the membership). The current draft has only 16 pages so space should not be an issue. 10 | 11 | 3. Clarify the difference between data and codata: both generally and in terms of difference it makes for main results of this paper. 12 | 13 | 4. Consider explaining how to recover the ordinary setting where a language is a set of words and language union is idempotent. For example, explain how the current work can be used to prove the classical theorem for the ordinary setting that the set of derivatives of a regular expression is finite modulo associativity, commutativity and idempotence of union (although the derivative defined here might require more). 14 | 15 | 5. Fix typos and other issues that have already been discussed in the reviews and the response. 16 |17 | 18 | My response: 19 | 20 | 1. I added the following paragraph to the beginning of section 7 (Related Work): 21 | 22 | > The shift from languages as sets of strings to type-level predicates (i.e., proof relevance or “parsing”) is akin to weighted automata [Schützenberger 1961; Droste and Kuske 2019] and more generally to semiring-based parsing [Chomsky and Schützenberger 1959; Goodman 1998, 1999; Liu 2004], noting that types (“Set” in Agda) form a (commutative) semiring (up to isomorphism). In particular, Lombardy and Sakarovitch [2005] investigated Brzozowski-style derivatives in this more general setting, formalizing and generalizing the work of Antimirov [1996], who decomposed derivatives into simpler components (“partial derivatives”) that sum to the full derivative and lead to efficient recognizers in the form of nondeterministic automata. 23 | 24 | Also an initial sentence to the (now) third paragraph of that section: 25 | 26 | > The main contributions of this paper are to formalization and mechanization of language recognition and parsing. 27 | 28 | 2. I added a new figure 2 with examples of languages and membership proofs, along with a new footnote 2 giving some detailed explanations. 29 | 30 | 3. I added a new footnote 10: 31 | 32 | > Inductive types (“data”) describe finitely large values and are processed recursively (as least fixed points) by pattern- matching clauses that decompose arguments. Dually, coinductive types (“codata”) describe infinitely large values and are processed corecursively (as greatest fixed points) by copattern-matching clauses that compose results. Programs on inductive types are often proved by induction, while programs on coinductive types are often proved by coinduction. (As we will see in Section 6.2, however, the simple relationship to the specification in Section 2 allows many trivial correctness proofs, needing neither induction or coinduction.) See Gordon [1995] for a tutorial on theory and techniques of coinduction. 33 | > 34 | > Tries in general represent functions, with each trie datum corresponding to a domain value. Even when each domain value is finitely large, there are often infinitely many of them (e.g., lists), so tries will naturally be infinite and thus more amenable to coinductive than inductive analysis. 35 | 36 | 4. Appended to the final paragraph of section 7: 37 | 38 | > As mentioned at the end of Section 3, one can simplify the development above somewhat by replacing proof isomorphism with the coarser notion of logical equivalence. Doing so restores union idempotence (modulo equivalence), losing proof relevance and thus re- placing parsing by recognition. One might then also prove that the set of derivatives of a regular language is finite (again, modulo equivalence). 39 | 40 | 5. Done! 41 | -------------------------------------------------------------------------------- /Predicate.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | {-# OPTIONS --safe --without-K #-} 3 | 4 | open import Level 5 | 6 | module Predicate {ℓ : Level} where 7 | 8 | open import Algebra.Core 9 | open import Data.Sum 10 | open import Data.Product 11 | open import Relation.Binary.PropositionalEquality hiding ([_]) 12 | open import Data.List 13 | open import Data.List.Relation.Unary.All 14 | 15 | open import Misc {ℓ} 16 | \end{code} 17 | 18 | %<*Pred> 19 | \AgdaTarget{Pred} 20 | \begin{code}[hide] 21 | Pred : Set ℓ → Set (suc ℓ) 22 | \end{code} 23 | \begin{code} 24 | Pred A = A → Set ℓ 25 | \end{code} 26 | % 27 | 28 | \begin{code} 29 | private 30 | variable 31 | A B C D : Set ℓ 32 | P Q R S : Pred A 33 | \end{code} 34 | 35 | Let's now generalize from predicates over lists (``languages'') to arbitrary predicates. 36 | 37 | First, we can transform types (predicate codomains) covariantly, with convenient nullary, unary, and binary variants: 38 | %<*codomain-transformers> 39 | \AgdaTarget{pureᵀ, mapᵀ, mapᵀ₂} 40 | \begin{code} 41 | pureᵀ : Set ℓ → Pred A 42 | pureᵀ x a = x 43 | 44 | mapᵀ : (Set ℓ → Set ℓ) → (Pred A → Pred A) 45 | mapᵀ g P a = g (P a) 46 | 47 | mapᵀ₂ : (Set ℓ → Set ℓ → Set ℓ) → 48 | (Pred A → Pred A → Pred A) 49 | mapᵀ₂ h P Q a = h (P a) (Q a) 50 | \end{code} 51 | % 52 | With these generalizations, we can easily define union and intersection with their identities (the empty and universal predicates), as well as complement:\footnote{All of these operations are standard \stdlibCite{Relation.Unary}.} 53 | %% \AgdaTarget{\_∪\_, ∪, ∅, \_∩\_, ∩, 𝒰, ∁} 54 | \begin{code}[hide] 55 | infixr 6 _∪_ 56 | infixr 7 _∩_ 57 | infixr 7 _·_ 58 | 59 | ∅ : Pred A 60 | 𝒰 : Pred A 61 | _∪_ : Op₂ (Pred A) 62 | _∩_ : Op₂ (Pred A) 63 | _·_ : Set ℓ → Op₁ (Pred A) 64 | ∁ : Op₁ (Pred A) 65 | \end{code} 66 | %<*codomain-ops> 67 | \begin{code} 68 | ∅ = pureᵀ ⊥ 69 | 𝒰 = pureᵀ ⊤ 70 | _∪_ = mapᵀ₂ _⊎_ 71 | _∩_ = mapᵀ₂ _×_ 72 | _·_ s = mapᵀ (s ×_) 73 | ∁ = mapᵀ ¬_ 74 | \end{code} 75 | % 76 | 77 | Next, transform \emph{values} (predicate domains) covariantly:\footnote{Predicate product (\AF{\_⟨×⟩\_}) is standard \stdlibCite{Relation.Unary}.}\notefoot{I think I'll have to rework the \AF{mapⱽ} definition when I generalize from types to other (semi)rings.} 78 | %<*domain-transformers> 79 | \AgdaTarget{pureⱽ, mapⱽ, mapⱽ₂} 80 | \begin{code} 81 | pureⱽ : A → Pred A 82 | pureⱽ x a = a ≡ x 83 | 84 | mapⱽ : (A → B) → (Pred A → Pred B) 85 | mapⱽ g P b = ∃ λ a → b ≡ g a × P a 86 | 87 | mapⱽ₂ : (A → B → C) → 88 | (Pred A → Pred B → Pred C) 89 | mapⱽ₂ h P Q c = ∃⇃ λ (a , b) → c ≡ h a b × P a × Q b 90 | \end{code} 91 | \begin{code}[hide] 92 | -- Alternatively 93 | 94 | infixr 2 _⟨×⟩_ 95 | _⟨×⟩_ : Pred A → Pred B → Pred (A × B) 96 | (P ⟨×⟩ Q) (a , b) = P a × Q b 97 | 98 | mapⱽ₂′ : (A → B → C) → (Pred A → Pred B → Pred C) 99 | mapⱽ₂′ g P Q = mapⱽ (uncurry g) (P ⟨×⟩ Q) 100 | \end{code} 101 | % 102 | Note that {\AF{mapⱽ} \AB{g} \AB{P}} is the image of the subset \AB{P} under the function \AB{g}. 103 | 104 | These domain transformations generalize concatenation and its identity to arbitrary binary operations or even operations of any arity. 105 | Rather than specialize all the way back to lists at this point, it will be useful to generalize to a binary operation \AB{\_∙\_} and an element \AB{ε}, which will form a monoid: 106 | % \AgdaTarget{MonoidOps, 𝟏, _⋆_, ⋆, \_☆, ☆, zero✪, suc✪} 107 | %<*domain-ops> 108 | \begin{AgdaSuppressSpace} 109 | \begin{code}[hide] 110 | module MonoidOps {M : Set ℓ} (_∙_ : Op₂ M) (ε : M) where 111 | 𝟏 : Pred M 112 | infixl 7 _⋆_ 113 | _⋆_ : Op₂ (Pred M) 114 | 115 | infixl 10 _☆ 116 | _☆ : Op₁ (Pred M) 117 | \end{code} 118 | \begin{code} 119 | 𝟏 = pureⱽ ε 120 | _⋆_ = mapⱽ₂ _∙_ 121 | P ☆ = mapⱽ (foldr _∙_ ε) (All P) 122 | \end{code} 123 | \begin{code}[hide] 124 | infixl 10 _✪ 125 | data _✪ (P : Pred M) : Pred M where 126 | zero✪ : (P ✪) ε 127 | suc✪ : ∀ {w} → (P ⋆ P ✪) w → (P ✪) w 128 | 129 | module ListOps (A : Set ℓ) where 130 | open MonoidOps {M = A ✶} _⊙_ [] public 131 | 132 | Lang : Set (suc ℓ) 133 | Lang = Pred (A ✶) 134 | 135 | ` : A → Lang 136 | \end{code} 137 | \begin{code} 138 | ` c = pureⱽ [ c ] 139 | \end{code} 140 | \end{AgdaSuppressSpace} 141 | % 142 | %<*Lang> 143 | \begin{code}[hide] 144 | private 145 | module Stuff where 146 | Lang⇃ : Set (suc ℓ) 147 | \end{code} 148 | \begin{code} 149 | Lang⇃ = Pred (A ✶) 150 | \end{code} 151 | % 152 | 153 | %% \begin{code} 154 | %% _✪ʳ : Op₁ (Pred M) 155 | %% data _✪ʳ P where 156 | %% zero✪ʳ : (P ✪ʳ) ε 157 | %% suc✪ʳ : ∀ {w} → (P ✪ʳ ⋆ P) w → (P ✪ʳ) w 158 | %% \end{code} 159 | -------------------------------------------------------------------------------- /To do.md: -------------------------------------------------------------------------------- 1 | ## To do 2 | 3 | * Conclusions and future work: 4 | * Sharing work and finite-state property 5 | * Check footnotes composition carefully. 6 | Sometimes there are terrible splits across pages or even a figure appearing beneath a footnote. 7 | The ACM layout does better than my own. 8 | 9 | * Mention optimizations for regular expressions and tries. 10 | * Read over `Predicate.Properties` and see what might be worth mentioning. 11 | For instance, distributivity of concatenation over union generalizes to mapping. 12 | * Rework abstract now that I've moved the predicate generalization toward the end of the paper. 13 | * Maybe further highlight the specific differences between the ν-δ lemmas and Brzozowski's original versions (imitated AFAICT by related work). 14 | 15 | * Maybe "proof relevance" is a better term for the proof isomorphism aspect of this paper. 16 | For languages, proof relevance means we care about the proofs, and we do care, since the parsing information is in the proof. 17 | "Proof-relevant languages". 18 | * Generalize from lists. 19 | Other types have their own deltas. 20 | 21 | Less important: 22 | 23 | * Fourier theorem? 24 | * Reconcile talk font vs paper font. 25 | I don't know how they come to differ. 26 | * Future work: 27 | * Beyond regular languages 28 | * Could I automate the transition from lemmas to implementation via rewrite rules? 29 | I'd have to suppress normalization, which would seem to break many proofs. 30 | * DFA and/or NFA implementations from the same specification and in the same style, i.e., indexed by type-level languages. 31 | 32 | From old paper version to-do: 33 | 34 | * Maybe switch to semiring and module notation starting in *Predicate Algebra*. 35 | * Explore the idea that Brzozowski's technique uses symbolic differentiation, but automatic differentiation has the usual advantages. 36 | What dimension do our domain and codomain have? 37 | Do forward, reverse, and mixed modes have anything to say? 38 | Is there something *linear* about derivatives? 39 | The right meaning for linearity seems clear from the function/free semimodule. 40 | * Other forms of convolution. 41 | I think generalized convolution is intrinsically amenable to memoization / dynamic programming. 42 | What about incremental evaluation? 43 | * Try redefining `AgdaFormat` as described in latex/agda.sty for more control over token rendering. 44 | If successful, stop using a special unicode replacement for "`*`" in the text. 45 | * Regular expressions and beyond: 46 | * Define an inductive type family of regular expressions and a semantic function. 47 | Provide a `ClosedSemiring` instance and prove that the semantic function is a closed semiring homomorphism. 48 | Use as an alternative implementation of matching. 49 | * Likewise for other language classes. 50 | * Run this stuff, and do some simple timing. 51 | * Find some languages with genuinely dependent types. 52 | I have something similar with decidability, though `Dec` instead of `Set`. 53 | 54 | ## Did 55 | 56 | * Link to repo, and say what versions of Agda and agda-stdlib. 57 | * From reviews: 58 | * Note versions of Agda and its standard library. 59 | * "line 283: I think it would be good to also say what is the correctness condition (in addition to the fact that it is guaranteed by types)." 60 | The condition is stated on line 228 and is captured in the type of `⟦_⟧` on line 310, namely that `⟦_⟧` is both computable and converts a language into a decidable recognizer/parser. 61 | Add a reminder at line 283. 62 | * "line 308: This left triangle looks like a map operation. Maybe a comment about why this has to be part of the syntax." 63 | 64 | Indeed, the left triangle does reify a map operation. 65 | The reason it must be part of the inductive representation is the same as the other constructors, namely so that the core lemmas (figure 2) translate into an implementation. 66 | Add a clarifying comment. 67 | 68 | * "Figures 4 and 5: there is no definition of the language (A* -> Set) denoted by an expression and thus no formal statement that it indeed denotes the language in its type." 69 | 70 | This missing semantic function maps every `p : Lang P` to `P`, regardless of the specifics of `p`. 71 | Add it. 72 | 73 | "line 623: By a choice of an equivalence relation we are quotienting the (encapsulated) representation. It would be good to give a bit more intuition here (which quotient is this)." 74 | 75 | This comment is closely related to the previous one about the denotation of the decidable language representations and has the same answer. 76 | Equivalence of representations is *semantic* equivalence, which is to say equivalence of the indices (type-level languages) that capture the semantics. 77 | Add this clarification. 78 | 79 | * Define language star via `mapⱽ` and `All`. 80 | * Check each use of `\\agda` to see if removing a preceding blank line improves indentation. 81 | * Point to the GitHub repo in the non-ACM version of the paper. 82 | * Maybe use acmsmall for the public/non-anonymous version of my paper, tweaking some features. 83 | Reconsider each use of `\\ifacm`. 84 | * Fix layout for `\\figref{nu-delta-codomain-lemmas}` in ACM version. 85 | It now uses horizontal layout, requiring `small` for ACM. 86 | I have plenty of room, so stack vertically instead. 87 | * Automata perspective. 88 | * Clarify tries as "automatic differentiation". 89 | * Mention the *non-idempotence* of this notion of languages, due to choice of equivalence relation. 90 | * Maybe mention logical equivalence vs proof isomorphism again in the Related Work section. 91 | * Wouter's paper 92 | * Mention OstermannJabs2018. 93 | * Maybe demote sections 4 (Reflections), 5 (Symbolic Differentiation), and 6 (Automatic Differentiation) into subsections of a new section called "From languages to parsing". 94 | Probably replace all occurrences of "`\\secreftwo{Symbolic Differentiation}{Automatic Differentiation}`". 95 | * Consider exposing level polymorphism in the talk & paper. 96 | It's there just under the surface. 97 | * Polish the abstract, especially the second/middle paragraph. 98 | * Fill in missing references and explanations flagged in the text. 99 | * Cite Martin Löf type theory. 100 | * Related work 101 | * Introduction, including clear statement of contributions 102 | * Mention that everything in the paper is verified by Agda. 103 | Here's what Andreas Abel wrote in *Equational Reasoning about Formal Languages in Coalgebraic Style*: 104 | "All the definitions, theorems, and proofs of this paper have been extracted from Agda code via an Agda to LaTeX translation and are, thus, guaranteed to be correct (assuming the consistency of Agda itself)." 105 | -------------------------------------------------------------------------------- /Algebra/Construct/Function.agda: -------------------------------------------------------------------------------- 1 | -- Some algebraic properties of functions 2 | 3 | {-# OPTIONS --without-K --safe #-} 4 | 5 | -- Following the style of Algebra.Construct.DirectProduct, except that our 6 | -- domain is a type rather than a setoid. 7 | 8 | open import Level using (Level; _⊔_) 9 | 10 | module Algebra.Construct.Function {a} (A : Set a) where 11 | 12 | open import Data.Product using (_,_; proj₁; proj₂) 13 | open import Algebra 14 | 15 | open import Closed 16 | 17 | private 18 | variable 19 | b ℓ : Level 20 | 21 | magma : Magma b ℓ → Magma (a ⊔ b) (a ⊔ ℓ) 22 | magma M = record 23 | { Carrier = A → Carrier 24 | ; _≈_ = λ f g → ∀ {x} → f x ≈ g x 25 | ; _∙_ = λ f g → λ x → f x ∙ g x 26 | ; isMagma = record 27 | { isEquivalence = record 28 | { refl = refl 29 | ; sym = λ f≡g → sym f≡g 30 | ; trans = λ f≡g g≡h → trans f≡g g≡h 31 | } 32 | ; ∙-cong = λ f≡h g≡i → ∙-cong f≡h g≡i 33 | } 34 | } where open Magma M 35 | 36 | semigroup : Semigroup b ℓ → Semigroup (a ⊔ b) (a ⊔ ℓ) 37 | semigroup G = record 38 | { isSemigroup = record 39 | { isMagma = Magma.isMagma (magma magma-G) 40 | ; assoc = λ f g h {x} → assoc (f x) (g x) (h x) 41 | -- ; assoc = λ _ _ _ → assoc _ _ _ -- also works throughout 42 | } 43 | } where open Semigroup G renaming (magma to magma-G) 44 | 45 | band : Band b ℓ → Band (a ⊔ b) _ 46 | band B = record 47 | { isBand = record 48 | { isSemigroup = Semigroup.isSemigroup (semigroup semigroup-B) 49 | ; idem = λ f {x} → idem (f x) 50 | } 51 | } where open Band B renaming (semigroup to semigroup-B) 52 | 53 | commutativeSemigroup : CommutativeSemigroup b ℓ → 54 | CommutativeSemigroup (a ⊔ b) _ 55 | commutativeSemigroup G = record 56 | { isCommutativeSemigroup = record 57 | { isSemigroup = Semigroup.isSemigroup (semigroup semigroup-G) 58 | ; comm = λ f g {x} → comm (f x) (g x) 59 | } 60 | } where open CommutativeSemigroup G renaming (semigroup to semigroup-G) 61 | 62 | semilattice : Semilattice b ℓ → Semilattice (a ⊔ b) _ 63 | semilattice L = record 64 | { isSemilattice = record 65 | { isBand = Band.isBand (band band-L) 66 | ; comm = λ f g {x} → comm (f x) (g x) 67 | } 68 | } where open Semilattice L renaming (band to band-L) 69 | 70 | monoid : Monoid b ℓ → Monoid (a ⊔ b) _ 71 | monoid M = record 72 | { ε = λ _ → ε 73 | ; isMonoid = record 74 | { isSemigroup = Semigroup.isSemigroup (semigroup semigroup-M) 75 | ; identity = (λ f {x} → identityˡ (f x)) , (λ f {x} → identityʳ (f x)) 76 | } 77 | } where open Monoid M renaming (semigroup to semigroup-M) 78 | 79 | commutativeMonoid : CommutativeMonoid b ℓ → CommutativeMonoid (a ⊔ b) _ 80 | commutativeMonoid M = record 81 | { isCommutativeMonoid = record 82 | { isMonoid = Monoid.isMonoid (monoid monoid-M) 83 | ; comm = λ f g {x} → comm (f x) (g x) 84 | } 85 | } where open CommutativeMonoid M renaming (monoid to monoid-M) 86 | 87 | -- TODO: IdempotentCommutativeMonoid, Group, AbelianGroup, Lattice, 88 | -- DistributiveLattice, NearSemiring, SemiringWithoutOne, 89 | -- CommutativeSemiringWithoutOne, SemiringWithoutOne, 90 | -- SemiringWithoutAnnihilatingZero 91 | 92 | semiringWithoutAnnihilatingZero : SemiringWithoutAnnihilatingZero b ℓ 93 | → SemiringWithoutAnnihilatingZero (b ⊔ a) (a ⊔ ℓ) 94 | semiringWithoutAnnihilatingZero R = record 95 | { isSemiringWithoutAnnihilatingZero = record 96 | { +-isCommutativeMonoid = CommutativeMonoid.isCommutativeMonoid 97 | (commutativeMonoid +-commutativeMonoid) 98 | ; *-isMonoid = Monoid.isMonoid (monoid (*-monoid)) 99 | ; distrib = (λ f g h {x} → distribˡ (f x) (g x) (h x)) 100 | , (λ f g h {x} → distribʳ (f x) (g x) (h x)) 101 | } 102 | } where open SemiringWithoutAnnihilatingZero R 103 | 104 | semiring : Semiring b ℓ → Semiring (a ⊔ b) _ 105 | semiring R = record 106 | { isSemiring = record 107 | { isSemiringWithoutAnnihilatingZero = 108 | SemiringWithoutAnnihilatingZero.isSemiringWithoutAnnihilatingZero 109 | (semiringWithoutAnnihilatingZero semiringWithoutAnnihilatingZero-R) 110 | ; zero = (λ f {x} → zeroˡ (f x)) , (λ f {x} → zeroʳ (f x)) } 111 | } where open Semiring R renaming 112 | (semiringWithoutAnnihilatingZero to semiringWithoutAnnihilatingZero-R) 113 | 114 | commutativeSemiring : CommutativeSemiring b ℓ → CommutativeSemiring (a ⊔ b) _ 115 | commutativeSemiring M = record 116 | { isCommutativeSemiring = record 117 | { isSemiring = Semiring.isSemiring (semiring semiring-M) 118 | ; *-comm = λ f g {x} → *-comm (f x) (g x) 119 | } 120 | } where open CommutativeSemiring M renaming (semiring to semiring-M) 121 | 122 | closedSemiring : ClosedSemiring b ℓ → ClosedSemiring (a ⊔ b) _ 123 | closedSemiring M = record 124 | { _✯ = λ f a → f a ✯ 125 | ; isClosedSemiring = record 126 | { isSemiring = Semiring.isSemiring (semiring semiring-M) 127 | ; star = λ f {x} → star (f x) 128 | } 129 | } where open ClosedSemiring M renaming (semiring to semiring-M) 130 | 131 | closedCommutativeSemiring : ClosedCommutativeSemiring b ℓ 132 | → ClosedCommutativeSemiring (a ⊔ b) _ 133 | closedCommutativeSemiring M = record 134 | { isClosedCommutativeSemiring = record 135 | { isCommutativeSemiring = CommutativeSemiring.isCommutativeSemiring 136 | (commutativeSemiring commutativeSemiring-M) 137 | ; star = λ f {x} → star (f x) 138 | } 139 | } where open ClosedCommutativeSemiring M renaming 140 | (commutativeSemiring to commutativeSemiring-M) 141 | 142 | 143 | ----- Module flavors, e.g., a module from a ring. 144 | 145 | -- We're going straight from semiring to semimodule here. Now I think a better 146 | -- design is to lift a semimodule, which might come from a semiring via 147 | -- Algebra.Module.Construct.TensorUnit. 148 | 149 | open import Algebra.Module.Bundles 150 | 151 | leftSemimodule : ∀ {r ℓr} {R : Semiring r ℓr} → LeftSemimodule R (a ⊔ r) (a ⊔ ℓr) 152 | leftSemimodule {R = R} = record 153 | { isLeftSemimodule = record 154 | { +ᴹ-isCommutativeMonoid = CommutativeMonoid.isCommutativeMonoid 155 | (commutativeMonoid +-commutativeMonoid) 156 | ; isPreleftSemimodule = record 157 | { *ₗ-cong = λ x~y f~g → *-cong x~y f~g 158 | ; *ₗ-zeroˡ = λ f {a} → proj₁ zero (f a) 159 | ; *ₗ-distribʳ = λ f s t {a} → distribʳ (f a) s t 160 | ; *ₗ-identityˡ = λ f {a} → *-identityˡ (f a) 161 | ; *ₗ-assoc = λ s t f {a} → *-assoc s t (f a) 162 | ; *ₗ-zeroʳ = λ s → zeroʳ s 163 | ; *ₗ-distribˡ = λ s f g {a} → distribˡ s (f a) (g a) 164 | } 165 | } 166 | } where open Semiring R 167 | 168 | rightSemimodule : ∀ {r ℓr} {R : Semiring r ℓr} → RightSemimodule R (a ⊔ r) (a ⊔ ℓr) 169 | rightSemimodule {R = R} = record 170 | { isRightSemimodule = record 171 | { +ᴹ-isCommutativeMonoid = CommutativeMonoid.isCommutativeMonoid 172 | (commutativeMonoid +-commutativeMonoid) 173 | ; isPrerightSemimodule = record 174 | { *ᵣ-cong = λ x~y f~g → *-cong x~y f~g 175 | ; *ᵣ-zeroʳ = λ f {a} → proj₂ zero (f a) 176 | ; *ᵣ-distribˡ = λ f s t {a} → distribˡ (f a) s t 177 | ; *ᵣ-identityʳ = λ f {a} → *-identityʳ (f a) 178 | ; *ᵣ-assoc = λ f s t {a} → *-assoc (f a) s t 179 | ; *ᵣ-zeroˡ = λ s → zeroˡ s 180 | ; *ᵣ-distribʳ = λ f g s {a} → distribʳ f (g a) (s a) 181 | } 182 | } 183 | } where open Semiring R 184 | 185 | leftModule : ∀ {r ℓr} {R : Ring r ℓr} → LeftModule R (a ⊔ r) (a ⊔ ℓr) 186 | leftModule {R = R} = record 187 | { isLeftModule = record 188 | { isLeftSemimodule = LeftSemimodule.isLeftSemimodule leftSemimodule 189 | ; -ᴹ‿cong = λ f~g {a} → -‿cong (f~g {a}) 190 | ; -ᴹ‿inverse = (λ f {a} → proj₁ -‿inverse (f a)) 191 | , (λ f {a} → proj₂ -‿inverse (f a)) 192 | } 193 | } where open Ring R 194 | 195 | semimodule : ∀ {r ℓr} {R : CommutativeSemiring r ℓr} → Semimodule R (a ⊔ r) (a ⊔ ℓr) 196 | semimodule {R = R} = record 197 | { isSemimodule = record 198 | { isBisemimodule = record 199 | { +ᴹ-isCommutativeMonoid = CommutativeMonoid.isCommutativeMonoid 200 | (commutativeMonoid +-commutativeMonoid) 201 | ; isPreleftSemimodule = 202 | LeftSemimodule.isPreleftSemimodule leftSemimodule 203 | ; isPrerightSemimodule = 204 | RightSemimodule.isPrerightSemimodule rightSemimodule 205 | ; *ₗ-*ᵣ-assoc = λ s f t {a} → *-assoc s (f a) t 206 | } 207 | } 208 | } where open CommutativeSemiring R 209 | 210 | 211 | -- TODO: more module flavors 212 | 213 | 214 | -------------------------------------------------------------------------------- /talk.tex: -------------------------------------------------------------------------------- 1 | %% Style copied from https://github.com/omelkonian/presentations/tree/master/%5B2019.08.20%5D%20BitML%20(SRC%20Presentation%20%40%20ICFP)) 2 | 3 | \newif\iftalk 4 | \newif\ifacm 5 | 6 | \talktrue 7 | 8 | \newif\ifshort 9 | 10 | \shorttrue 11 | 12 | %% \documentclass[aspectratio=43]{beamer} 13 | \documentclass[aspectratio=169]{beamer} 14 | 15 | \usetheme[ 16 | % block=fill, 17 | % background=light, 18 | % titleformat=smallcaps, 19 | % progressbar=frametitle, 20 | % numbering=none, 21 | ]{metropolis}%Warsaw,Madrid 22 | \setbeamersize{text margin left=.5cm,text margin right=.5cm} 23 | 24 | \usepackage{catchfilebetweentags} 25 | \usepackage[useregional]{datetime2} 26 | 27 | \RequirePackage{newunicodechar, amssymb, stmaryrd, unicode-math, setspace, comment} 28 | 29 | %% \renewcommand\alert[1]{\textcolor{mLightBrown}{#1}} 30 | 31 | \input{commands} 32 | \input{unicode} 33 | \input{macros} 34 | 35 | \usepackage{libertine} %% [tt=false] 36 | \usepackage{agda}% references 37 | 38 | \setmathfont{XITSMath-Regular.otf} 39 | \newunicodechar{◇}{\raisebox{0.23ex}{\ensuremath{_{\diamond}\hspace{-0.352em}}}} 40 | \newunicodechar{□}{\raisebox{-0.13ex}{\ensuremath{\scaleobj{0.3}{\Box}}\hspace{-0.23em}}} 41 | \newunicodechar{∷}{\ensuremath{\mathbin{\hspace{-0.2em}:\hspace{-0.27em}:\hspace{-0.15em}}}} 42 | 43 | %% \renewcommand{\AgdaFontStyle}[1]{\text{#1}} 44 | 45 | %---------------------------------------------------------------------------- 46 | 47 | \title{Symbolic and Automatic Differentiation of Languages} 48 | % \subtitle{...} 49 | \author{Conal Elliott} 50 | \date{ICFP, August 2021} 51 | %% \titlegraphic{ 52 | %% \vspace*{7cm} 53 | %% \includegraphics[keepaspectratio=true,height=1.4cm]{uu} 54 | %% \hspace{1cm} 55 | %% \includegraphics[keepaspectratio=true,height=1.4cm]{iohk} 56 | 57 | 58 | \begin{document} 59 | 60 | \begin{center} 61 | \setbeamerfont{title}{size=\large} 62 | \setbeamerfont{subtitle}{size=\small} 63 | \maketitle 64 | \setbeamerfont{title}{size=\Large} 65 | \setbeamerfont{subtitle}{size=\large} 66 | \end{center} 67 | 68 | \rnc\source{Language} 69 | 70 | \begin{frame}{Languages via type theory} 71 | \setstretch{0.8} 72 | \vspace{2.5ex} 73 | A language is a predicate on ``strings'' yielding proofs/explanations/parsings. 74 | \vspace{1ex} 75 | 76 | \agda{Lang-ops} 77 | 78 | \pause 79 | \emph{Puzzle:} how to compute parsings? 80 | \end{frame} 81 | 82 | \begin{frame}{Plan} 83 | \begin{itemize}\itemsep4ex 84 | \item 85 | Normal form for arbitrary (type-level) languages 86 | \item 87 | Normal form lemmas for standard language building blocks 88 | \item 89 | Decidability 90 | \item 91 | As corollaries, dual verified parsing implementations 92 | \end{itemize} 93 | \end{frame} 94 | 95 | \rnc\source{Calculus} 96 | 97 | \begin{frame}{Decomposing languages} 98 | \vspace{2ex} 99 | Consider each list constructor: 100 | \vspace{1.5ex} 101 | \agda{νδ} 102 | 103 | \pause 104 | \vspace{2ex} 105 | \AF{ν} and repeated \AF{δ} capture languages fully: 106 | 107 | \vspace{-1ex} 108 | \begin{minipage}[c]{2.8in} 109 | \agda{ν∘foldlδ} 110 | \end{minipage} 111 | \hfill 112 | \begin{minipage}[c]{2.5in} 113 | \mathindent0ex 114 | \agda{foldl} 115 | \end{minipage} 116 | \hfill\; 117 | \end{frame} 118 | 119 | \begin{frame}{Language calculus lemmas} 120 | \pause 121 | \vspace{-1ex} 122 | \vfill 123 | \agda{νδ-lemmas} 124 | \end{frame} 125 | 126 | \rnc\source{Decidability} 127 | 128 | \begin{frame}{Decidable types} 129 | \pause 130 | \vspace{2ex} 131 | \begin{minipage}{2in} 132 | \agda{Dec} 133 | \end{minipage} 134 | \vspace{1.5ex} 135 | \begin{minipage}{2in} 136 | \vspace{6.3ex} 137 | (\agda{¬}) 138 | \end{minipage} 139 | \hfill\; 140 | 141 | \pause 142 | \vspace{1ex} 143 | For predicates (languages), 144 | \agda{Decidable} 145 | 146 | \pause 147 | \vspace{2ex} 148 | Isomorphisms: 149 | 150 | \agda{isomorphisms} 151 | \end{frame} 152 | 153 | \begin{frame}{Compositionally decidable types} 154 | \pause 155 | \vspace{3ex} 156 | \agda{compositional-dec} 157 | %% Wait until I'm ready to talk about algebra 158 | %% Another closed semiring! 159 | \end{frame} 160 | 161 | \rnc\source{Reflections} 162 | 163 | \begin{frame}{Reflections} 164 | \pause 165 | \vspace{1.5ex} 166 | \agda{summary-a} 167 | \pause 168 | \agda{summary-b} 169 | 170 | \pause 171 | \setstretch{1.5} 172 | \AF{ν} and \AF{δ} are \emph{higher-order functions}. 173 | How to apply these rules? 174 | 175 | \pause 176 | Same challenge in differential calculus. 177 | 178 | Solve via \emph{symbolic} or \emph{automatic} differentiation. 179 | %% We to \emph{reflect} functions back to recipes. 180 | \end{frame} 181 | 182 | \rnc\source{Symbolic} 183 | 184 | %% \nc\zfill{\hfill} 185 | \nc\zfill{} 186 | 187 | \begin{frame}{Regular expressions \zfill (inductive)} 188 | \hypertarget{regex}{} 189 | \vspace{0.5ex} 190 | \setstretch{1.0} 191 | \agda{api} 192 | \end{frame} 193 | 194 | \begin{frame}{Symbolic differentiation \zfill (column-major / patterns)} 195 | \pause 196 | \hypertarget{symbolic-diff}{} 197 | \mathindent0ex 198 | \vspace{-2ex} 199 | \agda{defs} 200 | \end{frame} 201 | 202 | \rnc\source{Automatic} 203 | 204 | \begin{frame}{Tries \zfill \hyperlink{regex}{(coinductive)}} 205 | \pause 206 | \vspace{-1.5ex} 207 | \setstretch{1.0} 208 | \agda{api} 209 | \vspace{2ex} 210 | \end{frame} 211 | 212 | \agda{rules} 213 | 214 | \begin{frame}{Automatic differentiation \zfill (\hyperlink{symbolic-diff}{row-major / copatterns})} 215 | \pause 216 | \hypertarget{automatic-diff}{} 217 | \vspace{-1.75ex} 218 | \agda{defs} 219 | \end{frame} 220 | 221 | 222 | \ifshort 223 | 224 | \begin{frame}{Summary} 225 | \begin{itemize}\itemsep4ex 226 | \item 227 | Simple, non-computational formal specification. 228 | \item 229 | Reasoning from propositionally defined languages to decidable parsing via decision types. 230 | \item 231 | Duality of regular expressions and tries (symbolic and automatic differentiation).\\[2ex] 232 | Same code with dual interpretations. 233 | \end{itemize} 234 | \end{frame} 235 | 236 | \begin{frame}{Also in the paper} 237 | 238 | \begin{itemize}\itemsep4ex 239 | 240 | \item Termination checking 241 | \item Generalizing languages 242 | \item Algebraic properties of languages 243 | \item Transporting algebraic properties 244 | 245 | \end{itemize} 246 | \end{frame} 247 | 248 | \else 249 | 250 | \begin{frame}{Termination checking} 251 | Agda cannot prove termination for the {\AFi{δ} (\AB{p} \AF{⋆} \AB{q})} clause: 252 | \agda{termination} 253 | \vfill 254 | Andreas Abel solved this puzzle using sized types: 255 | \begin{center} 256 | \emph{\href{http://www.cse.chalmers.se/~abela/jlamp17.pdf}{Equational reasoning about formal languages in coalgebraic style}} (2016). 257 | \end{center} 258 | \end{frame} 259 | 260 | \rnc\source{SizedAutomatic} 261 | 262 | \begin{frame}{Sized tries \hfill (coinductive)} 263 | \vspace{-1.5ex} 264 | \setstretch{1.0} 265 | \agda{api} 266 | \vspace{2ex} 267 | \end{frame} 268 | 269 | \begin{frame}{Automatic differentiation (sized) \hfill (row-major \hyperlink{automatic-diff}{/} copatterns)} 270 | \vspace{-1.75ex} 271 | \agda{defs} 272 | \end{frame} 273 | 274 | 275 | \rnc\source{Predicate} 276 | 277 | \begin{comment} 278 | \begin{frame}{Two kinds of language operations} 279 | Codomain transformations: \AF{∅}, \AF 𝒰, \AF{\_∪\_}, \AF{\_∩\_} 280 | 281 | Domain transformations: \AF{𝟏}, \AF{\_⋆\_}, \AF{\_☆}, \AF{`} 282 | \end{frame} 283 | \end{comment} 284 | 285 | \begin{frame}{Generalizing languages to predicates} 286 | \vspace{-1ex} 287 | \begin{center} 288 | \mathindent-12ex 289 | \agda{Pred} 290 | \end{center} 291 | \vspace{-2ex} 292 | \hfill 293 | \mathindent0ex 294 | \small 295 | \begin{minipage}[b]{20em} 296 | \agda{codomain-transformers} 297 | \end{minipage} 298 | \hfill 299 | \begin{minipage}[b]{20em} 300 | \agda{domain-transformers} 301 | \end{minipage} 302 | \hfill\; 303 | \end{frame} 304 | 305 | \begin{frame}{Specializing predicates} 306 | \vspace{2ex} 307 | \mathindent0ex 308 | \begin{center} 309 | \mathindent-12ex 310 | \agda{Lang} 311 | \end{center} 312 | \vspace{-2ex} 313 | \hfill 314 | \begin{minipage}[c]{11em} 315 | \setstretch{1.7} 316 | \agda{codomain-ops} 317 | \end{minipage} 318 | \hfill 319 | \begin{minipage}[c]{15em} 320 | \setstretch{2.5} 321 | \agda{domain-ops} 322 | \end{minipage} 323 | \hfill\; 324 | \end{frame} 325 | 326 | \begin{frame}{Generalized ν and δ lemmas for codomain operations} 327 | \setstretch{1.7} 328 | \centering 329 | \ExecuteMetaData[Calculus.tex]{νδ-codomain} 330 | \end{frame} 331 | 332 | \begin{frame}{Algebraic properties of languages} 333 | \begin{itemize}\itemsep4ex 334 | \item Predicates form a semimodule and closed semiring. 335 | \item Functions from monoids to semirings (e.g., languages) form an additional semiring. 336 | \item Commutative if the monoid and semiring are. 337 | \item Algebraic abstractions are simply typed.\\[2ex] 338 | Wrap indexed types: {\AF{∃} \AD{Dec}}, {\AF{∃} \AD{(\ARe{Lang} \APo ∞)}}, etc. 339 | Index-equivalence. 340 | \end{itemize} 341 | \end{frame} 342 | 343 | \rnc\source{Transport} 344 | 345 | \begin{frame}{Transporting algebraic properties} 346 | \vspace{1.5ex} 347 | \AgdaEmptySkip2ex 348 | \agda{examples} 349 | \end{frame} 350 | 351 | \fi 352 | 353 | \end{document} 354 | -------------------------------------------------------------------------------- /unicode.tex: -------------------------------------------------------------------------------- 1 | \RequirePackage{bbding} 2 | 3 | %% \usepackage[utf8]{inputenc} %% ignored with xelatex 4 | %% \usepackage{amssymb} 5 | 6 | %%%% From https://github.com/gallais/nary 7 | 8 | %% % Misc symbols 9 | %% \newunicodechar{′}{\ensuremath{\prime}} 10 | %% \newunicodechar{−}{\ensuremath{-}} 11 | %% \newunicodechar{─}{\ensuremath{-}} 12 | %% \newunicodechar{◆}{\ensuremath{\Diamondblack}} 13 | %% \newunicodechar{⧫}{\ensuremath{\blacklozenge}} 14 | \newunicodechar{∷}{\ensuremath{\mathbin{::}}} 15 | \newunicodechar{∙}{\ensuremath{\bullet}} 16 | %% \newunicodechar{□}{\ensuremath{\Box}} 17 | %% \newunicodechar{∎}{\ensuremath{\blacksquare}} 18 | %% \newunicodechar{⋆}{\ensuremath{\star}} 19 | 20 | % indices 21 | \newunicodechar{₀}{\ensuremath{_0}} 22 | \newunicodechar{₁}{\ensuremath{_1}} 23 | \newunicodechar{₂}{\ensuremath{_2}} 24 | \newunicodechar{₃}{\ensuremath{_3}} 25 | 26 | %% \newunicodechar{ₖ}{\ensuremath{_k}} 27 | %% \newunicodechar{ₘ}{\ensuremath{_m}} 28 | \newunicodechar{ₙ}{\ensuremath{_n}} 29 | %% \newunicodechar{ᵣ}{\ensuremath{_r}} 30 | %% \newunicodechar{ₛ}{\ensuremath{_s}} 31 | 32 | %% \newunicodechar{₊}{\ensuremath{_+}} 33 | 34 | %% % exponents 35 | %% \newunicodechar{⁺}{\ensuremath{\textsuperscript{+}}} 36 | %% \newunicodechar{⁻}{\ensuremath{\textsuperscript{-}}} 37 | 38 | %% \newunicodechar{²}{\ensuremath{^2}} 39 | 40 | \newunicodechar{ⁱ}{\ensuremath{^i}} 41 | \newunicodechar{ˡ}{\raisebox{-2pt}{\ensuremath{^l}}} 42 | \newunicodechar{ʳ}{\raisebox{-1pt}{\ensuremath{^r}}} 43 | 44 | %% \newunicodechar{ᴬ}{\ensuremath{^A}} 45 | %% \newunicodechar{ᴮ}{\ensuremath{^B}} 46 | %% \newunicodechar{ᴵ}{\ensuremath{^I}} 47 | %% \newunicodechar{ᴿ}{\ensuremath{^R}} 48 | %% \newunicodechar{ᵀ}{\ensuremath{^T}} 49 | %% \newunicodechar{ᵁ}{\ensuremath{^U}} 50 | %% \newunicodechar{ⱽ}{\ensuremath{^V}} 51 | 52 | %% % Dots 53 | %% \newunicodechar{⋯}{\ensuremath{\cdots}} 54 | 55 | %% % Equality symbols 56 | \newunicodechar{≡}{\ensuremath{\equiv}} 57 | %% \newunicodechar{≢}{\ensuremath{\not\equiv}} 58 | \newunicodechar{≟}{\mbox{\tiny\ensuremath{\stackrel{?}{=}}}} 59 | \newunicodechar{≈}{\ensuremath{\approx}} 60 | 61 | %% % Ordering symbols 62 | %% \newunicodechar{≤}{\ensuremath{\le}} 63 | %% \newunicodechar{≥}{\ensuremath{\ge}} 64 | 65 | %% % Arrows 66 | \newunicodechar{→}{\ensuremath{\rightarrow}} 67 | %% \newunicodechar{←}{\ensuremath{\leftarrow}} 68 | %% \newunicodechar{⇒}{\ensuremath{\Rightarrow}} 69 | %% \newunicodechar{⇉}{\ensuremath{\rightrightarrows}} 70 | 71 | %% % Mathematical symbols 72 | %% \newunicodechar{∂}{\ensuremath{\partial}} 73 | %% \newunicodechar{∋}{\ensuremath{\ni}} 74 | %% \newunicodechar{∞}{\ensuremath{\infty}} 75 | \newunicodechar{∀}{\ensuremath{\forall}} 76 | \newunicodechar{∃}{\ensuremath{\exists}} 77 | %% \newunicodechar{⊢}{\ensuremath{\vdash}} 78 | \newunicodechar{⟨}{\ensuremath{\langle}} 79 | \newunicodechar{⟩}{\ensuremath{\rangle}} 80 | %% \newunicodechar{⊤}{\ensuremath{\top}} 81 | \newunicodechar{∘}{\ensuremath{\circ}} 82 | \newunicodechar{⊎}{\ensuremath{\uplus}} 83 | \newunicodechar{×}{\ensuremath{\times}} 84 | %% \newunicodechar{ℕ}{\ensuremath{\mathbb{N}}} 85 | %% \newunicodechar{⟦}{\ensuremath{\llbracket}} 86 | %% \newunicodechar{⟧}{\ensuremath{\rrbracket}} 87 | \newunicodechar{∈}{\ensuremath{\in}} 88 | %% \newunicodechar{↑}{\ensuremath{\uparrow}} 89 | \newunicodechar{¬}{\ensuremath{\neg}} 90 | \newunicodechar{⊥}{\ensuremath{\bot}} 91 | %% \newunicodechar{↝}{\ensuremath{\leadsto}} 92 | %% \newunicodechar{↶}{\ensuremath{\curvearrowleft}} 93 | %% \newunicodechar{↺}{\ensuremath{\circlearrowleft}} 94 | \newunicodechar{⊔}{\ensuremath{\sqcup}} 95 | %% \newunicodechar{⨆}{\ensuremath{\bigsqcup}} 96 | \newunicodechar{∩}{\ensuremath{\cap}} 97 | \newunicodechar{∪}{\ensuremath{\cup}} 98 | 99 | 100 | %%%%%%%%%%%%%%%%% LETTERS 101 | 102 | %% % Misc 103 | \newunicodechar{ℓ}{\ensuremath{\ell}} 104 | 105 | %% % Greek uppercase 106 | %% \newunicodechar{Δ}{\ensuremath{\Delta}} 107 | %% \newunicodechar{Γ}{\ensuremath{\Gamma}} 108 | %% \newunicodechar{Σ}{\ensuremath{\Sigma}} 109 | %% \newunicodechar{Θ}{\ensuremath{\Theta}} 110 | %% \newunicodechar{Ω}{\ensuremath{\Omega}} 111 | 112 | %% % Greek lowercase 113 | \newunicodechar{α}{\ensuremath{\alpha}} 114 | \newunicodechar{β}{\ensuremath{\beta}} 115 | \newunicodechar{γ}{\ensuremath{\gamma}} 116 | \newunicodechar{δ}{\ensuremath{\delta}} 117 | \newunicodechar{ε}{\ensuremath{\varepsilon}} 118 | \newunicodechar{φ}{\ensuremath{\phi}} 119 | \newunicodechar{γ}{\ensuremath{\gamma}} 120 | \newunicodechar{ι}{\ensuremath{\iota}} 121 | \newunicodechar{κ}{\ensuremath{\kappa}} 122 | \newunicodechar{λ}{\ensuremath{\lambda}} 123 | \newunicodechar{μ}{\ensuremath{\mu}} 124 | \newunicodechar{ψ}{\ensuremath{\psi}} 125 | \newunicodechar{η}{\ensuremath{\eta}} 126 | \newunicodechar{ρ}{\ensuremath{\rho}} 127 | \newunicodechar{σ}{\ensuremath{\sigma}} 128 | \newunicodechar{τ}{\ensuremath{\tau}} 129 | \newunicodechar{ξ}{\ensuremath{\xi}} 130 | \newunicodechar{ζ}{\ensuremath{\zeta}} 131 | \newunicodechar{Π}{\ensuremath{\Pi}} 132 | 133 | 134 | %% % mathcal 135 | %% \newunicodechar{𝓒}{\ensuremath{\mathcal{C}}} 136 | %% \newunicodechar{𝓔}{\ensuremath{\mathcal{E}}} 137 | %% \newunicodechar{𝓕}{\ensuremath{\mathcal{F}}} 138 | %% \newunicodechar{𝓡}{\ensuremath{\mathcal{R}}} 139 | %% \newunicodechar{𝓢}{\ensuremath{\mathcal{S}}} 140 | %% \newunicodechar{𝓣}{\ensuremath{\mathcal{T}}} 141 | %% \newunicodechar{𝓥}{\ensuremath{\mathcal{V}}} 142 | %% \newunicodechar{𝓦}{\ensuremath{\mathcal{W}}} 143 | 144 | \newunicodechar{𝒟}{\ensuremath{\mathscr{D}}} 145 | 146 | %% Above I copied from gallais's TyDe'19 paper. 147 | %% Below are my additions. 148 | 149 | %%https://tex.stackexchange.com/questions/486120/using-the-unicode-prime-character-outside-math-mode-with-unicode-math-and-newuni 150 | \AtBeginDocument{\newunicodechar{′}{\ensuremath{'}}} 151 | 152 | \newunicodechar{ν}{\ensuremath{\nu}} 153 | \newunicodechar{∨}{\ensuremath{\vee}} 154 | \newunicodechar{∧}{\ensuremath{\wedge}} 155 | 156 | % \newunicodechar{∅}{\ensuremath{\emptyset}} 157 | \newunicodechar{∅}{\ensuremath{\varnothing}} 158 | \newunicodechar{∁}{\ensuremath{\complement}} 159 | 160 | \newunicodechar{ᵒ}{\ensuremath{^o}} 161 | 162 | \newunicodechar{𝒰}{\ensuremath{\mathscr{U}}} 163 | 164 | \newunicodechar{𝟏}{\ensuremath{\mathbf{1}}} 165 | 166 | \newcommand\smallOp[2]{\ensuremath{\mathbin{\mbox{\raisebox{#1}{\scriptsize $#2$}}}}} 167 | \iftalk 168 | \newcommand\append{\smallOp{0.5pt}{+\!\!+}} 169 | \else 170 | \newcommand\append{\mathbin{+\!\!+}} 171 | \fi 172 | 173 | \newunicodechar{⊙}{\ensuremath{\append}} 174 | 175 | %% \newunicodechar{ℓ}{\ensuremath{\mathcal{l}}} 176 | 177 | \newcommand\Star[1]{\ensuremath{\mathop{\hspace{-0.0em}\raisebox{0.6ex}{\tiny #1}}}} 178 | 179 | \newunicodechar{✯}{\Star \FiveStarOutlineHeavy} 180 | \newunicodechar{☆}{\Star \FiveStarShadow} 181 | \newunicodechar{✪}{\Star \FiveStarShadow} 182 | \newunicodechar{✰}{\Star \FiveStarShadow} 183 | \newunicodechar{✶}{\Star \SixStar} 184 | \newunicodechar{✦}{\Star \FourStarOpen} 185 | \newunicodechar{✸}{\Star \EightStarBold} 186 | 187 | %% \newunicodechar{⋆}{\smallOp{0.75pt}{}} 188 | %% \newunicodechar{⋆}{\raisebox{-2pt}{\Star \Asterisk}} 189 | \newunicodechar{⋆}{\ensuremath{\ast}} 190 | %% \newunicodechar{⋆}{\ensuremath{*}}% "*" seems fine in Libertine. 191 | 192 | \newunicodechar{✲}{\raisebox{-0.8ex}{\hspace{-0.17em}*}} 193 | 194 | \newunicodechar{⇔}{\ensuremath{\Leftrightarrow}} 195 | \newunicodechar{↔}{\ensuremath{\leftrightarrow}} 196 | \newunicodechar{⟷}{\ensuremath{\longleftrightarrow}} 197 | 198 | \newunicodechar{⊤}{{\small \ensuremath{\top}}} 199 | 200 | %% I haven't gotten amsmath to work, so use \Box instead of \blacksquare for now. 201 | \newunicodechar{∎}{\ensuremath{\blacksquare}} 202 | %% \newunicodechar{∎}{\ensuremath{\Box}} 203 | 204 | \newunicodechar{𝟘}{\ensuremath{\mathbb{0}}} 205 | \newunicodechar{𝟙}{\ensuremath{\mathbb{1}}} 206 | 207 | \newunicodechar{⁻}{\ensuremath{^-}} 208 | \newunicodechar{⁰}{\ensuremath{^0}} 209 | \newunicodechar{¹}{\ensuremath{^1}} 210 | \newunicodechar{²}{\ensuremath{^2}} 211 | %% \newunicodechar{ʸ}{\ensuremath{^y}} 212 | %% \newunicodechar{ⁿ}{\ensuremath{^{\hspace{0.3pt}n}}} 213 | \newunicodechar{ᴸ}{\ensuremath{^L}} 214 | 215 | \newunicodechar{ᴹ}{\ensuremath{^{\!M\!}}} 216 | 217 | %% \newunicodechar{ᵗ}{\ensuremath{^T}} 218 | 219 | %% \newunicodechar{ᵀ}{\ensuremath{^τ}} 220 | %% \newunicodechar{ⱽ}{\ensuremath{^v}} 221 | \newunicodechar{ᵀ}{\ensuremath{^o}} 222 | \newunicodechar{ⱽ}{\ensuremath{^i}} 223 | 224 | \newunicodechar{ₗ}{\ensuremath{_{\raisebox{2pt}{\tiny \!$l$}}}} 225 | \newunicodechar{ᵣ}{\ensuremath{_r}} 226 | 227 | \newunicodechar{∼}{\ensuremath{\sim}} 228 | \newunicodechar{⌊}{\ensuremath{\lfloor}} 229 | \newunicodechar{⌋}{\ensuremath{\rfloor}} 230 | \newunicodechar{≗}{\ensuremath{\circeq}} 231 | 232 | \newunicodechar{◃}{\raisebox{0.2ex}{\tiny \ensuremath{\triangleleft}}} 233 | \newunicodechar{▹}{\raisebox{0.2ex}{\tiny \ensuremath{\triangleright}}} 234 | \newunicodechar{◂}{\raisebox{0.2ex}{\tiny \ensuremath{\blacktriangleleft}}} 235 | \newunicodechar{▸}{\raisebox{0.2ex}{\tiny \ensuremath{\blacktriangleright}}} 236 | 237 | \newunicodechar{·}{\ensuremath{\cdot}} 238 | 239 | %% Invisible characters for otherwise clashing definitions. 240 | \newunicodechar{⇂}{} 241 | \newunicodechar{⇃}{} 242 | 243 | %% Unobtrusive module name. Decorate the following dot. Note vertical in ex 244 | %% and horizontal in em to adjust with font size. 245 | %% \newunicodechar{◇}{\raisebox{0.25ex}{\ensuremath{_{\diamond}\hspace{-0.34em}}}} 246 | 247 | %% Another good trick: cause the next symbol to be hat'd, math style, e.g., ʻν ʻδ. 248 | %% A variation on this trick can replace a module dot, etc. 249 | \newcommand\mathHat[1]{\ensuremath{\hat{\mbox{#1}}}} 250 | \newunicodechar{ʻ}{\mathHat} 251 | 252 | \newunicodechar{‽}{\ensuremath{^?}} 253 | \newunicodechar{″}{\hspace{1pt}′} 254 | -------------------------------------------------------------------------------- /talk.md: -------------------------------------------------------------------------------- 1 | # ICFP talk 2 | 3 | I generally don't write out my talks, but this time I did, since I'm recording it. 4 | 5 | ## Languages via type theory 6 | 7 | Languages are often formalized as sets of strings. 8 | Alternatively, we can consider a language to be a type-level predicate, that is a function that maps any string to a type of membership proofs. 9 | Each inhabitant of a membership type is an explanation or parsing. 10 | If there are no inhabitants, then the string is not in the language. 11 | 12 | With this simple idea, we can easily define the usual building blocks of languages. 13 | *[Reveal.]* 14 | 15 | For instance, a string is in the union of `P` and `Q` when it's in `P` or in `Q`, so memberships are sum types, corresponding to logical disjunction. 16 | Similarly, membership proofs for language intersections are products, corresponding to logical conjunction. 17 | 18 | As a more interesting example, a string `w` is in the concatenation of languages `P` and `Q`, exactly when there are strings in `P` and `Q` that concatenate to `w`. 19 | While precise and elegant, this definition does not tell us how to recognize or parse language concatenations, since the existential quantification is over infinitely many string pairs. 20 | 21 | In case you're not used to Agda, let me point out a few things: 22 | 23 | * Bottom and top are the empty and singleton types. 24 | Logically, they correspond to falsity and truth. 25 | * The type `x ≡ y` has a single inhabitant when indeed `x` does equal `y` and otherwise is empty. 26 | * Sum types are written with disjoint union symbol. 27 | * Non-dependent product types are written with cross symbol. 28 | * Dependent products types are written with the exists symbol. 29 | 30 | *[Reveal.]* 31 | The puzzle addressed by this paper is how to bridge the gap between this non-computational specification and correct, computable parsing. 32 | 33 | ## Plan 34 | 35 | Here's the plan: 36 | 37 | * First, we'll define a normal form for languages. 38 | The main idea is to use Brzozowski-style language derivatives at the level of types, which is to say propositions. 39 | 40 | * Next, prove lemmas relating this normal form to the usual building blocks of languages. 41 | 42 | * So far, everything is propositional rather than computable, that is at the level of types rather than computational inhabitants. 43 | The next step relates our language lemmas to decidable form. 44 | 45 | * Finally, we'll apply these insights in two ways to yield dual correct parsing implementations. 46 | 47 | ## Decomposing languages 48 | 49 | Remembering that a language is a function from lists, let's consider each data constructor for lists, namely nil and cons: 50 | 51 | * Nullability of a language is the proposition that the language contains the empty string. 52 | * The derivative of a language `P` with respect to a leading character `a` is another language, consisting of the tails of the strings in `P` that start with `a`. 53 | The proofs that `w` is in `δ P a` are exactly the proofs that `a ∷ w` is in `P`. 54 | 55 | The importance of these definitions comes from a simple fact with a simple inductive proof. 56 | Given a language `P` and string, we can successively differentiate `P` with respect to characters in the string, resulting in a residual language. 57 | The original language `P` contains the input string exactly if the residual language is nullable. 58 | 59 | This theorem and its proof are very simple, as shown in the lower left of the slide. 60 | Agda is quite liberal with names, and in this case, the name of the theorem is "`ν∘foldlδ`" with no spaces. 61 | 62 | Note that everything on this slide applies not just to languages but more generally to functions from lists to anything. 63 | 64 | ## Language calculus 65 | 66 | The next step in our plan to identify and prove a collection of lemmas that relate ν and δ to the standard language building blocks. 67 | 68 | *[reveal]* 69 | 70 | The ν lemmas on the left relate types and are equalities or isomorphisms. 71 | The δ lemmas on the right relate languages and are equalities or extensional isomorphisms. 72 | 73 | Agda proves the equalities automatically, simply by normalization. 74 | The others take a bit of effort. 75 | All proofs are in the paper's source code. 76 | 77 | The style of these lemmas is important. 78 | Each one reduces ν or δ of a standard language construction to ν and/or δ of simpler constructions. 79 | The computable implementations that follow and their full correctness (including termination) are corollaries of these lemmas. 80 | 81 | ## Decidable types 82 | 83 | Given a language and a candidate string, we can apply the language to the string to yield a type of proofs that the string is in the language. 84 | Now we want to *construct* such a proof or show that one cannot exist. 85 | We can express this goal as a decision data type. 86 | 87 | A parser for a language is then a computable function that maps an arbitrary string to a decision about the string belonging to the language. 88 | 89 | *[reveal]* 90 | 91 | Isomorphisms appear in the language lemmas of the previous slide, so we'll need to know how they relate to decidability. 92 | Fortunately, the answers are simple. 93 | If two types are isomorphic, then a decision for one suffices to decide the other, since evidence of each can be mapped to evidence of the other. 94 | Likewise for extensionally isomorphic predicates. 95 | 96 | ## Compositionally decidable types 97 | 98 | While we cannot possibly decide *all* predicates, we can decide some of them, and we can do so compositionally. 99 | Together with the isomorphism deciders from the previous slide, these compositional deciders cover all of the constructions appearing in the language calculus lemmas. 100 | For instance, consider deciding the conjunction of `A` and `B`. 101 | If we have proofs of each, we can pair those proofs to form a proof of the conjunction. 102 | On the other hand, if we have a *disproof* of either, we can use it to construct a disproof of the conjunction. 103 | 104 | ## Reflections 105 | 106 | Let's pause to reflect on where we are in the story. 107 | 108 | * The language decomposition theorem reduces membership to a succession of derivatives, followed by nullability. 109 | * Our language lemmas tell us how to decompose nullability and derivatives of languages to the same questions on simpler languages. 110 | * The resulting questions are all expressed in terms of propositions and predicates that happen to be compositionally decidable. 111 | 112 | It looks like we're done: we just formulate the languages, compute derivatives and nullability, and apply the language lemmas. 113 | 114 | However, we cannot simply apply the language lemmas by pattern matching, because languages are functions and so cannot be inspected structurally. 115 | 116 | Exactly this same situation holds in differential calculus, since differentiation in that setting is also defined on functions and we have a collection of lemmas about derivatives for various formulations of functions, such as sums, products, trig functions, and compositions. 117 | When we want to *compute* correct derivatives, there are two standard solutions, known as "symbolic" and "automatic" differentiation: 118 | 119 | * Symbolic differentiation represents functions structually and applies the differentiation rules by pattern matching. 120 | 121 | * Automatic differentiation represents differentiable functions as functions that compute their derivatives as well as their regular values. 122 | 123 | We can apply these same strategies to languages. 124 | 125 | ## Regular expressions (inductive) 126 | 127 | Applying the first strategy to languages leads to an inductive data type of regular expressions. 128 | To keep a simple and rigorous connection to our original specification, we'll index this inductive data type by the type-level language it denotes. 129 | Here I've kept the vocabulary the same, distinguishing type-level language building blocks by a module prefix that shows here as a small, lowered diamond. 130 | 131 | Note that ν and δ here are decidable versions. 132 | Correctness of parsing is guaranteed by the types of these two functions, so any definitions that type-check will suffice. 133 | 134 | ## Symbolic differentiation (column-major / patterns) 135 | 136 | Given our inductive representation, we need only define ν and δ. 137 | The definitions shown on this slide are systematically derived from the language lemmas shown earlier in the talk, using the compositional deciders. 138 | The definitions are to be read in column-major order, that is, each column is one definition, given by pattern matching. 139 | 140 | ## Tries (coinductive) 141 | 142 | Applying the second strategy to languages leads instead to a *coinductive* type of tries, which is an exact dual to the inductive structure of regular expressions. 143 | The language operations are now defined functions instead of constructors, while ν and δ are now selectors instead of functions. 144 | Again, correctness of parsing is guaranteed by the types, so any definitions that type-check will suffice. 145 | 146 | ## Automatic differentiation (row-major / copatterns) 147 | 148 | Given our coinductive representation, we need only define the language building blocks. 149 | The definitions shown on this slide are *also* systematically derived from the language lemmas shown earlier in the talk, using the compositional deciders. 150 | This time, the definitions are to be read in row-major order, that is, each row is one definition, and is given by co-pattern matching. 151 | Remember that this language representation is a pair of a ν and a δ, so a single definition gives both. 152 | 153 | Here's the magical thing. 154 | Compare this implementation with the previous one. 155 | The code is exactly the same, but its interpretation changes, as signified by the syntax coloring, which the Agda compiler generates during typesetting. 156 | 157 | ## Summary 158 | 159 | What pleases me most about this work: 160 | 161 | * First, there is the simple formal specification, uncomplicated by computability. 162 | 163 | * Second, the clear path of reasoning from propositionally defined languages to decidable parsing via decision types. 164 | 165 | * Finally, the duality of regular expressions and tries---paralleling to symbolic and automatic differentiation---using exactly the same code for parsing in both cases, though with dual interpretations. 166 | 167 | ## Also in the paper 168 | 169 | There are more goodies in the paper. 170 | I encourage you to give it a read. 171 | -------------------------------------------------------------------------------- /Calculus.lagda: -------------------------------------------------------------------------------- 1 | \begin{code}[hide] 2 | {-# OPTIONS --safe #-} 3 | open import Level 4 | 5 | module Calculus {ℓ} (A : Set ℓ) where 6 | 7 | -- open import Data.Empty 8 | -- open import Data.Unit 9 | open import Data.Sum 10 | open import Data.Product 11 | open import Data.List 12 | open import Data.List.Properties using (++-identityʳ) 13 | open import Function using (id; _∘_; _↔_; mk↔′) 14 | open import Relation.Binary.PropositionalEquality hiding ([_]) 15 | 16 | open import Misc {ℓ} 17 | open import Inverses {ℓ} 18 | 19 | -- open import Language A 20 | open import Predicate ; open ListOps A 21 | 22 | private 23 | variable 24 | a c : A 25 | s x : Set ℓ 26 | b : Level 27 | B : Set b 28 | X : Set b 29 | P Q R : Lang 30 | f : A ✶ → B 31 | u v w : A ✶ 32 | \end{code} 33 | 34 | %<*νδ> 35 | \AgdaTarget{ν, δ} 36 | \begin{code} 37 | ν⇃ : Lang → Set ℓ -- “nullable” 38 | ν⇃ P = P [] 39 | 40 | δ⇃ : Lang → A → Lang -- “derivative” 41 | δ⇃ P a w = P (a ∷ w) 42 | \end{code} 43 | %νδ> 44 | 45 | %<*ν𝒟> 46 | \AgdaTarget{ν, δ} 47 | \begin{code} 48 | ν : (A ✶ → B) → B -- “nullable” 49 | ν f = f [] 50 | 51 | 𝒟 : (A ✶ → B) → A ✶ → (A ✶ → B) -- “derivative” 52 | 𝒟 f u = λ v → f (u ⊙ v) 53 | \end{code} 54 | %% 𝒟 f u v = f (u ⊙ v) 55 | %% 𝒟 f u = λ v → f (u ⊙ v) 56 | %% 𝒟 f u = f ∘ (u ⊙_) 57 | %ν𝒟> 58 | 59 | %<*δ> 60 | \begin{code} 61 | δ : (A ✶ → B) → A → (A ✶ → B) 62 | δ f a = 𝒟 f [ a ] 63 | \end{code} 64 | %δ> 65 | 66 | %<*𝒟[]⊙> 67 | \begin{code} 68 | 𝒟[] : 𝒟 f [] ≡ f 69 | 70 | 𝒟⊙ : 𝒟 f (u ⊙ v) ≡ 𝒟 (𝒟 f u) v 71 | \end{code} 72 | \begin{code}[hide] 73 | 𝒟[] = refl 74 | 75 | 𝒟⊙ {u = []} = refl 76 | 𝒟⊙ {f = f} {u = a ∷ u} = 𝒟⊙ {f = δ f a} {u = u} 77 | \end{code} 78 | %𝒟[]⊙> 79 | 80 | %<*foldl> 81 | \begin{code}[hide] 82 | private 83 | module Stuff where 84 | \end{code} 85 | \begin{code} 86 | foldl⇃ : (X → A → X) → X → A ✶ → X 87 | foldl⇃ h x [] = x 88 | foldl⇃ h x (a ∷ as) = foldl⇃ h (h x a) as 89 | \end{code} 90 | % 91 | 92 | %<*ν∘𝒟> 93 | \begin{code} 94 | ν∘𝒟 : ν ∘ 𝒟 f ≗ f 95 | \end{code} 96 | \begin{code}[hide] 97 | ν∘𝒟 u rewrite (++-identityʳ u) = refl 98 | 99 | -- ν∘𝒟 f u = cong f (++-identityʳ u) 100 | 101 | -- ν∘𝒟 f [] = refl 102 | -- ν∘𝒟 f (a ∷ as) = ν∘𝒟 (δ f a) as 103 | \end{code} 104 | %ν∘𝒟> 105 | 106 | %<*𝒟foldl> 107 | \begin{code} 108 | 𝒟foldl : 𝒟 f ≗ foldl δ f 109 | \end{code} 110 | \begin{code}[hide] 111 | 𝒟foldl [] = refl 112 | 𝒟foldl (a ∷ as) = 𝒟foldl as 113 | \end{code} 114 | %𝒟foldl> 115 | 116 | %% Phasing out. Still used in talk.tex. 117 | %<*ν∘foldlδ> 118 | \AgdaTarget{ν∘foldlδ} 119 | %% ν∘foldlδ : ∀ w → P w ≡ ν (foldl δ P w) 120 | \begin{code} 121 | ν∘foldlδ : ν ∘ foldl δ P ≗ P 122 | ν∘foldlδ [] = refl 123 | ν∘foldlδ (a ∷ as) = ν∘foldlδ as 124 | \end{code} 125 | %ν∘foldlδ> 126 | 127 | \begin{code} 128 | open import Algebra.Core 129 | private 130 | variable 131 | g : Op₁ (Set ℓ) 132 | h : Op₂ (Set ℓ) 133 | \end{code} 134 | 135 | %<*νδ-codomain> 136 | \AgdaTarget{νpureᵀ, νmapᵀ, νmapᵀ₂} 137 | \begin{code} 138 | νpureᵀ : ν (pureᵀ x) ≡ x 139 | νmapᵀ : ν (mapᵀ g P) ≡ g (ν P) 140 | νmapᵀ₂ : ν (mapᵀ₂ h P Q) ≡ h (ν P) (ν Q) 141 | \end{code} 142 | \AgdaTarget{δpureᵀ, δmapᵀ, δmapᵀ₂} 143 | \begin{code} 144 | δpureᵀ : δ (pureᵀ x) a ≡ pureᵀ x 145 | δmapᵀ : δ (mapᵀ g P) a ≡ mapᵀ g (δ P a) 146 | δmapᵀ₂ : δ (mapᵀ₂ h P Q) a ≡ mapᵀ₂ h (δ P a) (δ Q a) 147 | \end{code} 148 | \begin{code}[hide] 149 | νpureᵀ = refl 150 | δpureᵀ = refl 151 | νmapᵀ = refl 152 | δmapᵀ = refl 153 | νmapᵀ₂ = refl 154 | δmapᵀ₂ = refl 155 | \end{code} 156 | %νδ-codomain> 157 | 158 | %<*νδ-lemmas> 159 | {\mathindent0ex 160 | \hfill 161 | \setstretch{1.7} 162 | \begin{minipage}{28ex} 163 | \AgdaTarget{ν∅, ν∪, ν𝟏, ν⋆, ν✪, δ∅, δ∪, δ𝟏, δ⋆, δ✪} 164 | \begin{code} 165 | ν∅ : ν ∅ ≡ ⊥ 166 | ν𝒰 : ν 𝒰 ≡ ⊤ 167 | ν∪ : ν (P ∪ Q) ≡ (ν P ⊎ ν Q) 168 | ν∩ : ν (P ∩ Q) ≡ (ν P × ν Q) 169 | ν· : ν (s · P) ≡ (s × ν P) 170 | ν𝟏 : ν 𝟏 ↔ ⊤ 171 | ν⋆ : ν (P ⋆ Q) ↔ (ν P × ν Q) 172 | ν☆ : ν (P ☆) ↔ (ν P) ✶ 173 | ν` : ν (` c) ↔ ⊥ 174 | \end{code} 175 | \end{minipage} 176 | \hfill 177 | \begin{minipage}{38ex} 178 | \begin{code} 179 | δ∅ : δ ∅ a ≡ ∅ 180 | δ𝒰 : δ 𝒰 a ≡ 𝒰 181 | δ∪ : δ (P ∪ Q) a ≡ δ P a ∪ δ Q a 182 | δ∩ : δ (P ∩ Q) a ≡ δ P a ∩ δ Q a 183 | δ· : δ (s · P) a ≡ s · δ P a 184 | δ𝟏 : δ 𝟏 a ⟷ ∅ 185 | δ⋆ : δ (P ⋆ Q) a ⟷ ν P · δ Q a ∪ δ P a ⋆ Q 186 | δ☆ : δ (P ☆) a ⟷ (ν P) ✶ · (δ P a ⋆ P ☆) 187 | δ` : δ (` c) a ⟷ (a ≡ c) · 𝟏 188 | \end{code} 189 | \end{minipage} 190 | \hfill\; 191 | } 192 | %νδ-lemmas> 193 | 194 | \begin{code} 195 | ν✪ : ν (P ✪) ↔ (ν P) ✶ 196 | δ✪ : δ (P ✪) a ⟷ (ν P) ✶ · (δ P a ⋆ P ✪) 197 | \end{code} 198 | 199 | \begin{code}[hide] 200 | 201 | δ· = refl 202 | δ∅ = refl 203 | δ∩ = refl 204 | δ∪ = refl 205 | δ𝒰 = refl 206 | ν· = refl 207 | ν∅ = refl 208 | ν∩ = refl 209 | ν∪ = refl 210 | ν𝒰 = refl 211 | 212 | ν𝟏 = mk↔′ 213 | (λ { refl → tt }) 214 | (λ { tt → refl }) 215 | (λ { tt → refl }) 216 | (λ { refl → refl }) 217 | 218 | δ𝟏 = mk↔′ (λ ()) (λ ()) (λ ()) (λ ()) 219 | 220 | ν` = mk↔′ (λ ()) (λ ()) (λ ()) (λ ()) 221 | 222 | δ` = mk↔′ 223 | (λ { refl → refl , refl }) 224 | (λ { (refl , refl) → refl }) 225 | (λ { (refl , refl) → refl }) 226 | (λ { refl → refl }) 227 | 228 | ν⋆ = mk↔′ 229 | (λ { (([] , []) , refl , νP , νQ) → νP , νQ }) 230 | (λ { (νP , νQ) → ([] , []) , refl , νP , νQ }) 231 | (λ { (νP , νQ) → refl } ) 232 | (λ { (([] , []) , refl , νP , νQ) → refl}) 233 | 234 | δ⋆ {a = a} {w = w} = mk↔′ 235 | (λ { (([] , .(a ∷ w)) , refl , νP , Qaw) → inj₁ (νP , Qaw) 236 | ; ((.a ∷ u , v) , refl , Pu , Qv) → inj₂ ((u , v) , refl , Pu , Qv) }) 237 | (λ { (inj₁ (νP , Qaw)) → ([] , a ∷ w) , refl , νP , Qaw 238 | ; (inj₂ ((u , v) , refl , Pu , Qv)) → ((a ∷ u , v) , refl , Pu , Qv) }) 239 | (λ { (inj₁ (νP , Qaw)) → refl 240 | ; (inj₂ ((u , v) , refl , Pu , Qv)) → refl }) 241 | (λ { (([] , .(a ∷ w)) , refl , νP , Qaw) → refl 242 | ; ((.a ∷ u , v) , refl , Pu , Qv) → refl }) 243 | 244 | ν✪ {P = P} = mk↔′ k k⁻¹ invˡ invʳ 245 | where 246 | k : ν (P ✪) → (ν P) ✶ 247 | k zero✪ = [] 248 | k (suc✪ (([] , []) , refl , (νP , νP✪))) = νP ∷ k νP✪ 249 | 250 | k⁻¹ : (ν P) ✶ → ν (P ✪) 251 | k⁻¹ [] = zero✪ 252 | k⁻¹ (νP ∷ νP✶) = suc✪ (([] , []) , refl , (νP , k⁻¹ νP✶)) 253 | 254 | invˡ : ∀ (νP✶ : (ν P) ✶) → k (k⁻¹ νP✶) ≡ νP✶ 255 | invˡ [] = refl 256 | invˡ (νP ∷ νP✶) rewrite invˡ νP✶ = refl 257 | 258 | invʳ : ∀ (νP✪ : ν (P ✪)) → k⁻¹ (k νP✪) ≡ νP✪ 259 | invʳ zero✪ = refl 260 | invʳ (suc✪ (([] , []) , refl , (νP , νP✪))) rewrite invʳ νP✪ = refl 261 | 262 | δ✪ {P}{a} {w} = mk↔′ k k⁻¹ invˡ invʳ 263 | where 264 | k : δ (P ✪) a w → ((ν P) ✶ · (δ P a ⋆ P ✪)) w 265 | k (suc✪ (([] , .(a ∷ w)) , refl , (νP , P✪a∷w))) with k P✪a∷w 266 | ... | νP✶ , etc 267 | = νP ∷ νP✶ , etc 268 | k (suc✪ ((.a ∷ u , v) , refl , (Pa∷u , P✪v))) = [] , (u , v) , refl , (Pa∷u , P✪v) 269 | 270 | k⁻¹ : ((ν P) ✶ · (δ P a ⋆ P ✪)) w → δ (P ✪) a w 271 | k⁻¹ ([] , (u , v) , refl , (Pa∷u , P✪v)) = (suc✪ ((a ∷ u , v) , refl , (Pa∷u , P✪v))) 272 | k⁻¹ (νP ∷ νP✶ , etc) = (suc✪ (([] , a ∷ w) , refl , (νP , k⁻¹ (νP✶ , etc)))) 273 | 274 | invˡ : (s : ((ν P) ✶ · (δ P a ⋆ P ✪)) w) → k (k⁻¹ s) ≡ s 275 | invˡ ([] , (u , v) , refl , (Pa∷u , P✪v)) = refl 276 | invˡ (νP ∷ νP✶ , etc) rewrite invˡ (νP✶ , etc) = refl 277 | 278 | invʳ : (s : δ (P ✪) a w) → k⁻¹ (k s) ≡ s 279 | invʳ (suc✪ (([] , .(a ∷ w)) , refl , (νP , P✪a∷w))) rewrite invʳ P✪a∷w = refl 280 | invʳ (suc✪ ((.a ∷ u , v) , refl , (Pa∷u , P✪v))) = refl 281 | 282 | \end{code} 283 | 284 | 285 | Clarify the relationship with automatic differentiation: 286 | 287 | Now enhance \AF 𝒟: 288 | %<*𝒟′> 289 | \begin{code} 290 | 𝒟′ : (A ✶ → B) → A ✶ → B × (A ✶ → B) 291 | 𝒟′ f u = f u , 𝒟 f u 292 | \end{code} 293 | %𝒟′> 294 | 295 | %% The ʻ name trick (defined in unicode.tex) adds a hat to the name it precedes. 296 | %<*ʻ𝒟> 297 | \begin{code} 298 | ʻ𝒟 : (A ✶ → B) → A ✶ → B × (A ✶ → B) 299 | ʻ𝒟 f u = let f″ = foldl δ f u in ν f″ , f″ 300 | \end{code} 301 | %% ʻ𝒟 f u = ν f″ , f″ where f″ = foldl δ f u 302 | %ʻ𝒟> 303 | %% Oddly, variables introduced in "where" don't look like variables 304 | 305 | %<*𝒟′≡ʻ𝒟> 306 | \begin{code} 307 | 𝒟′≡ʻ𝒟 : 𝒟′ f ≗ ʻ𝒟 f 308 | 𝒟′≡ʻ𝒟 [] = refl 309 | 𝒟′≡ʻ𝒟 (a ∷ as) = 𝒟′≡ʻ𝒟 as 310 | \end{code} 311 | %𝒟′≡ʻ𝒟> 312 | 313 | \note{There's probably a way to \AK{rewrite} with \AB{𝒟foldl} and \AB{𝒟′foldl} to eliminate the induction here.} 314 | 315 | 316 | Experiment with alternative star: 317 | \begin{code} 318 | open import Data.List.Relation.Unary.All 319 | open import Data.List.Properties 320 | open import Data.Sum.Algebra 321 | open import Data.Product.Algebra using (×-cong) 322 | open import Predicate.Properties 323 | open MonoidSemiringProperties (++-isMonoid {A = A}) 324 | open import Closed.Instances ; open Types {ℓ} 325 | \end{code} 326 | 327 | \begin{code} 328 | νfix : ν (P ☆) ↔ (⊤ ⊎ ν P × ν (P ☆)) 329 | νfix {P = P} = 330 | begin 331 | ν (P ☆) 332 | ≈⟨ ☆-star ⟩ 333 | (𝟏 ∪ P ⋆ P ☆) [] 334 | ≈⟨ ⊎-cong ν𝟏 ν⋆ ⟩ 335 | (⊤ ⊎ ν P × ν (P ☆)) 336 | ∎ where open ↔R 337 | 338 | δfix : δ (P ☆) a ⟷ (ν P · δ (P ☆) a ∪ δ P a ⋆ P ☆) 339 | δfix {P = P} {a = a} {w} = 340 | begin 341 | δ (P ☆) a w 342 | ≈⟨ ☆-star ⟩ 343 | (𝟏 ∪ P ⋆ P ☆) (a ∷ w) 344 | ≡⟨⟩ 345 | (𝟏 (a ∷ w) ⊎ (P ⋆ P ☆) (a ∷ w)) 346 | ≡⟨⟩ 347 | (δ 𝟏 a w ⊎ δ (P ⋆ P ☆) a w) 348 | ≈⟨ ⊎-cong δ𝟏 δ⋆ ⟩ 349 | ( ⊥ ⊎ (ν P · δ (P ☆) a ∪ δ P a ⋆ P ☆) w ) 350 | ≈⟨ ⊎-identityˡ ℓ _ ⟩ 351 | (ν P · δ (P ☆) a ∪ δ P a ⋆ P ☆) w 352 | -- ≡⟨⟩ 353 | -- (ν P × δ (P ☆) a w ⊎ (δ P a ⋆ P ☆) w) 354 | ∎ where open ↔R 355 | 356 | ν☆ {P = P} = 357 | begin 358 | ν (P ☆) 359 | ≈˘⟨ ✪↔☆ ⟩ 360 | ν (P ✪) 361 | ≈⟨ ν✪ ⟩ 362 | (ν P) ✶ 363 | ∎ where open ↔R 364 | 365 | δ☆ {P = P}{a} {w} = 366 | begin 367 | δ (P ☆) a w 368 | ≈˘⟨ ✪↔☆ ⟩ 369 | δ (P ✪) a w 370 | ≈⟨ δ✪ ⟩ 371 | ((ν P) ✶ · (δ P a ⋆ P ✪)) w 372 | ≈⟨ ×-congˡ (⋆-congˡ ✪↔☆) ⟩ 373 | ((ν P) ✶ · (δ P a ⋆ P ☆)) w 374 | ∎ where open ↔R 375 | 376 | -- TODO: keep looking for direct proofs of ν☆ and δ☆ so I can abandon ✪. 377 | \end{code} 378 | 379 | 380 | {- 381 | ν☆ {P = P} = mk↔′ k k⁻¹ invˡ invʳ 382 | where 383 | k : ν (P ☆) → (ν P) ✶ 384 | k (.[] , refl , []) = [] 385 | k ([] ∷ ws , eq , νp ∷ pws) = νp ∷ k (ws , eq , pws) 386 | 387 | k⁻¹ : (ν P) ✶ → ν (P ☆) 388 | k⁻¹ [] = [] , refl , [] 389 | k⁻¹ (νp ∷ νps) with k⁻¹ νps 390 | ... | ws , eq , pws = [] ∷ ws , eq , νp ∷ pws 391 | 392 | invˡ : ∀ (νP✶ : (ν P) ✶) → k (k⁻¹ νP✶) ≡ νP✶ 393 | invˡ [] = refl 394 | invˡ (νp ∷ νps) rewrite invˡ νps = refl 395 | 396 | invʳ : ∀ (νP☆ : ν (P ☆)) → k⁻¹ (k νP☆) ≡ νP☆ 397 | invʳ ([] , refl , []) = refl 398 | invʳ ([] ∷ ws , eq , νp ∷ pws) rewrite invʳ (ws , eq , pws) = refl 399 | 400 | δ☆ {P}{a} {w} = mk↔′ k k⁻¹ invˡ invʳ 401 | where 402 | k : δ (P ☆) a w → ((ν P) ✶ · (δ P a ⋆ P ☆)) w 403 | 404 | k ([] ∷ us , eq , νp ∷ pus) with k (us , eq , pus) 405 | ... | νps , etc = νp ∷ νps , etc 406 | k ((.a ∷ u) ∷ us , refl , pau ∷ pus) = 407 | [] , (u , foldr _⊙_ [] us) , refl , pau , us , refl , pus 408 | 409 | -- k⁻¹ : ((ν P) ✶ · (δ P a ⋆ P ✪)) w → δ (P ✪) a w 410 | -- k⁻¹ ([] , (u , v) , refl , (Pa∷u , P✪v)) = (suc✪ ((a ∷ u , v) , refl , (Pa∷u , P✪v))) 411 | -- k⁻¹ (νP ∷ νP✶ , etc) = (suc✪ (([] , a ∷ w) , refl , (νP , k⁻¹ (νP✶ , etc)))) 412 | 413 | k⁻¹ : ((ν P) ✶ · (δ P a ⋆ P ☆)) w → δ (P ☆) a w 414 | k⁻¹ ([] , (u , v) , refl , pau , psv) = {!!} , {!!} , {!!} 415 | k⁻¹ (x ∷ u , snd) = {!!} 416 | 417 | invˡ : (s : ((ν P) ✶ · (δ P a ⋆ P ☆)) w) → k (k⁻¹ s) ≡ s 418 | invˡ z = {!!} 419 | 420 | invʳ : (s : δ (P ☆) a w) → k⁻¹ (k s) ≡ s 421 | invʳ z = {!!} 422 | 423 | 424 | -- P ☆ = mapⱽ (foldr _∙_ ε) (All P) 425 | -} 426 | 427 | \end{code} 428 | -------------------------------------------------------------------------------- /Existential.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | 3 | -- Algebraic bundles for existentially quantified types, with equivalence 4 | -- defined as equivalence on proj₁. Status: coverage for all bundles in 5 | -- Algebra.Bundles and none in Algebra.Module.Bundles. 6 | 7 | {-# OPTIONS --safe --without-K #-} 8 | open import Level 9 | 10 | module Existential where 11 | 12 | open import Algebra 13 | open import Data.Product 14 | open import Function using (_on_) 15 | import Relation.Binary.Construct.On as On 16 | 17 | private variable b c ℓ : Level 18 | 19 | -- Lift operations to dependent products 20 | \end{code} 21 | 22 | %<*Dop> 23 | \begin{AgdaAlign} 24 | \begin{code}[hide] 25 | module Dop {A : Set c} (F : A → Set ℓ) where 26 | D₀ : A → Set ℓ 27 | D₁ : Op₁ A → Set (c ⊔ ℓ) 28 | D₂ : Op₂ A → Set (c ⊔ ℓ) 29 | \end{code} 30 | \begin{code} 31 | D₀ s = F s 32 | D₁ g = ∀ {a} → F a → F (g a) 33 | D₂ h = ∀ {a b} → F a → F b → F (h a b) 34 | \end{code} 35 | \end{AgdaAlign} 36 | % 37 | 38 | An equivalent, less factored version for the paper: 39 | %<*inj> 40 | \begin{code}[hide] 41 | module InjPaper {A : Set c} {F : A → Set ℓ} where 42 | open Dop F 43 | \end{code} 44 | \begin{code} 45 | inj₀ : ∀ {s : A} (s′ : F s) → ∃ F 46 | inj₀ {s} s′ = s , s′ 47 | 48 | inj₁ : ∀ {g : A → A} (g′ : ∀ {a} → F a → F (g a)) → (∃ F → ∃ F) 49 | inj₁ {g} g′ (a , x) = g a , g′ x 50 | 51 | inj₂ : ∀ {h : A → A → A} (h′ : ∀ {a b} → F a → F b → F (h a b)) → (∃ F → ∃ F → ∃ F) 52 | inj₂ {h} h′ (a , x) (b , y) = h a b , h′ x y 53 | \end{code} 54 | % 55 | 56 | \begin{code} 57 | module Inj {A : Set c} {F : A → Set ℓ} where 58 | 59 | open Dop F 60 | 61 | -- Mostly work with 2nd projections, ignoring and inferring 1st projections. 62 | pattern inj b = (_ , b) 63 | 64 | inj₁ : ∀ {∙_ : Op₁ A} (∙′_ : D₁ ∙_) → Op₁ (∃ F) 65 | inj₁ ∙′_ (inj x) = inj (∙′ x) 66 | 67 | inj₂ : ∀ {_∙_ : Op₂ A} (_∙′_ : D₂ _∙_) → Op₂ (∃ F) 68 | inj₂ _∙′_ (inj x) (inj y) = inj (x ∙′ y) 69 | \end{code} 70 | 71 | %<*prop> 72 | \begin{code}[hide] 73 | private 74 | variable 75 | h : Level 76 | P : A → Set h 77 | Q : A → A → Set h 78 | R : A → A → A → Set h 79 | \end{code} 80 | \begin{code} 81 | prop₁ : (∀ a → P a) → ∀ ((a , _) : ∃ F) → P a 82 | prop₁ P (a , _) = P a 83 | 84 | prop₂ : (∀ a b → Q a b) → ∀ ((a , _) (b , _) : ∃ F) → Q a b 85 | prop₂ Q (a , _) (b , _) = Q a b 86 | 87 | prop₃ : (∀ a b c → R a b c) → ∀ ((a , _) (b , _) (c , _) : ∃ F) → R a b c 88 | prop₃ R (a , _) (b , _) (c , _) = R a b c 89 | \end{code} 90 | % 91 | 92 | %<*prop₃′> 93 | \begin{code}[hide] 94 | prop₃′ : (∀ a b c → R a b c) → ∀ ((a , _) (b , _) (c , _) : ∃ F) → R a b c 95 | \end{code} 96 | \begin{code}[inline] 97 | prop₃′ R _ _ _ = R _ _ _ 98 | \end{code} 99 | % 100 | 101 | \begin{code} 102 | -- -- The propᵢ definitions can be mostly inferred, e.g., 103 | -- prop₃ P _ _ _ = P _ _ _ 104 | 105 | open Inj 106 | 107 | module _ (m : Magma c ℓ) (F : Magma.Carrier m → Set b) where 108 | open Magma m ; open Dop F 109 | mkMagma : D₂ _∙_ → Magma (c ⊔ b) ℓ 110 | mkMagma _∙′_ = record 111 | { Carrier = ∃ F 112 | ; _≈_ = _≈_ on proj₁ 113 | ; _∙_ = inj₂ _∙′_ 114 | ; isMagma = record 115 | { isEquivalence = On.isEquivalence proj₁ isEquivalence 116 | ; ∙-cong = ∙-cong 117 | } 118 | } 119 | 120 | module _ (g : Semigroup c ℓ) (F : Semigroup.Carrier g → Set b) where 121 | open Semigroup g hiding (isMagma) ; open Dop F 122 | mkSemigroup : D₂ _∙_ → Semigroup (c ⊔ b) ℓ 123 | mkSemigroup _∙′_ = record 124 | { isSemigroup = record 125 | { isMagma = isMagma 126 | ; assoc = prop₃ assoc 127 | } 128 | } where open Magma (mkMagma magma F _∙′_) 129 | 130 | module _ (g : Monoid c ℓ) (F : Monoid.Carrier g → Set b) where 131 | open Monoid g hiding (isSemigroup) ; open Dop F 132 | mkMonoid : D₂ _∙_ → D₀ ε → Monoid (c ⊔ b) ℓ 133 | mkMonoid _∙′_ ε′ = record 134 | { ε = inj ε′ 135 | ; isMonoid = record { isSemigroup = isSemigroup 136 | ; identity = prop₁ identityˡ , prop₁ identityʳ } 137 | } where open Semigroup (mkSemigroup semigroup F _∙′_) 138 | 139 | module _ (g : CommutativeMonoid c ℓ) (F : CommutativeMonoid.Carrier g → Set b) where 140 | open CommutativeMonoid g hiding (isMonoid) ; open Dop F 141 | mkCommutativeMonoid : D₂ _∙_ → D₀ ε → CommutativeMonoid (c ⊔ b) ℓ 142 | mkCommutativeMonoid _∙′_ ε′ = record 143 | { isCommutativeMonoid = record { isMonoid = isMonoid ; comm = prop₂ comm } 144 | } where open Monoid (mkMonoid monoid F _∙′_ ε′) 145 | 146 | module _ (g : Group c ℓ) (F : Group.Carrier g → Set b) where 147 | open Group g hiding (isMonoid) ; open Dop F 148 | mkGroup : D₂ _∙_ → D₀ ε → D₁ _⁻¹ → Group (c ⊔ b) ℓ 149 | mkGroup _∙′_ ε′ _⁻¹′ = record 150 | { _⁻¹ = inj₁ _⁻¹′ 151 | ; isGroup = record { isMonoid = isMonoid 152 | ; inverse = prop₁ inverseˡ , prop₁ inverseʳ 153 | ; ⁻¹-cong = ⁻¹-cong } 154 | } where open Monoid (mkMonoid monoid F _∙′_ ε′) 155 | 156 | module _ (g : AbelianGroup c ℓ) (F : AbelianGroup.Carrier g → Set b) where 157 | open AbelianGroup g hiding (isGroup) ; open Dop F 158 | mkAbelianGroup : D₂ _∙_ → D₀ ε → D₁ _⁻¹ → AbelianGroup (c ⊔ b) ℓ 159 | mkAbelianGroup _∙′_ ε′ _⁻¹′ = record 160 | { isAbelianGroup = record { isGroup = isGroup ; comm = prop₂ comm } 161 | } where open Group (mkGroup group F _∙′_ ε′ _⁻¹′) 162 | 163 | module _ (l : Lattice c ℓ) (F : Lattice.Carrier l → Set b) where 164 | open Lattice l hiding () ; open Dop F 165 | mkLattice : D₂ _∨_ → D₂ _∧_ → Lattice (c ⊔ b) ℓ 166 | mkLattice _∨′_ _∧′_ = record 167 | { Carrier = ∃ F 168 | ; _≈_ = _≈_ on proj₁ 169 | ; _∨_ = inj₂ _∨′_ 170 | ; _∧_ = inj₂ _∧′_ 171 | ; isLattice = record 172 | { isEquivalence = On.isEquivalence proj₁ isEquivalence 173 | ; ∨-comm = prop₂ ∨-comm 174 | ; ∨-assoc = prop₃ ∨-assoc 175 | ; ∨-cong = ∨-cong 176 | ; ∧-comm = prop₂ ∧-comm 177 | ; ∧-assoc = prop₃ ∧-assoc 178 | ; ∧-cong = ∧-cong 179 | ; absorptive = prop₂ (proj₁ absorptive) , prop₂ (proj₂ absorptive) 180 | } 181 | } 182 | 183 | module _ (g : DistributiveLattice c ℓ) (F : DistributiveLattice.Carrier g → Set b) where 184 | open DistributiveLattice g hiding (isLattice) ; open Dop F 185 | mkDistributiveLattice : D₂ _∨_ → D₂ _∧_ → DistributiveLattice (c ⊔ b) ℓ 186 | mkDistributiveLattice _∨′_ _∧′_ = record 187 | { isDistributiveLattice = record 188 | { isLattice = isLattice 189 | ; ∨-distribʳ-∧ = prop₃ ∨-distribʳ-∧ 190 | } 191 | } where open Lattice (mkLattice lattice F _∨′_ _∧′_) 192 | 193 | module _ (r : NearSemiring c ℓ) (F : NearSemiring.Carrier r → Set b) where 194 | open NearSemiring r hiding (+-isMonoid; *-isSemigroup) ; open Dop F 195 | mkNearSemiring : D₂ _+_ → D₂ _*_ → D₀ 0# → NearSemiring (c ⊔ b) ℓ 196 | mkNearSemiring _+′_ _*′_ 0#′ = record 197 | { isNearSemiring = record 198 | { +-isMonoid = +-isMonoid 199 | ; *-isSemigroup = *-isSemigroup 200 | ; distribʳ = prop₃ distribʳ 201 | ; zeroˡ = prop₁ zeroˡ 202 | } 203 | } where open Monoid (mkMonoid +-monoid F _+′_ 0#′) 204 | renaming (isMonoid to +-isMonoid) 205 | open Semigroup (mkSemigroup *-semigroup F _*′_) 206 | renaming (isSemigroup to *-isSemigroup) 207 | 208 | module _ (r : SemiringWithoutOne c ℓ) (F : SemiringWithoutOne.Carrier r → Set b) where 209 | open SemiringWithoutOne r hiding (+-isCommutativeMonoid; *-isSemigroup) ; open Dop F 210 | mkSemiringWithoutOne : D₂ _+_ → D₂ _*_ → D₀ 0# → SemiringWithoutOne (c ⊔ b) ℓ 211 | mkSemiringWithoutOne _+′_ _*′_ 0#′ = record 212 | { isSemiringWithoutOne = record 213 | { +-isCommutativeMonoid = +-isCommutativeMonoid 214 | ; *-isSemigroup = *-isSemigroup 215 | ; distrib = prop₃ (proj₁ distrib) , prop₃ (proj₂ distrib) 216 | ; zero = prop₁ zeroˡ , prop₁ zeroʳ 217 | } 218 | } where open CommutativeMonoid 219 | (mkCommutativeMonoid +-commutativeMonoid F _+′_ 0#′) 220 | renaming (isCommutativeMonoid to +-isCommutativeMonoid) 221 | open Semigroup (mkSemigroup *-semigroup F _*′_) 222 | renaming (isSemigroup to *-isSemigroup) 223 | 224 | module _ (r : CommutativeSemiringWithoutOne c ℓ) 225 | (F : CommutativeSemiringWithoutOne.Carrier r → Set b) where 226 | open CommutativeSemiringWithoutOne r hiding (isSemiringWithoutOne) 227 | open Dop F 228 | mkCommutativeSemiringWithoutOne : D₂ _+_ → D₂ _*_ → D₀ 0# 229 | → CommutativeSemiringWithoutOne (c ⊔ b) ℓ 230 | mkCommutativeSemiringWithoutOne _+′_ _*′_ 0#′ = record 231 | { isCommutativeSemiringWithoutOne = record 232 | { isSemiringWithoutOne = isSemiringWithoutOne 233 | ; *-comm = prop₂ *-comm 234 | } 235 | } where open SemiringWithoutOne 236 | (mkSemiringWithoutOne semiringWithoutOne F _+′_ _*′_ 0#′) 237 | 238 | module _ (r : SemiringWithoutAnnihilatingZero c ℓ) 239 | (F : SemiringWithoutAnnihilatingZero.Carrier r → Set b) where 240 | open SemiringWithoutAnnihilatingZero r hiding (+-isCommutativeMonoid; *-isMonoid) 241 | open Dop F 242 | mkSemiringWithoutAnnihilatingZero : D₂ _+_ → D₂ _*_ → D₀ 0# → D₀ 1# 243 | → SemiringWithoutAnnihilatingZero (c ⊔ b) ℓ 244 | mkSemiringWithoutAnnihilatingZero _+′_ _*′_ 0#′ 1#′ = record 245 | { isSemiringWithoutAnnihilatingZero = record 246 | { +-isCommutativeMonoid = +-isCommutativeMonoid 247 | ; *-isMonoid = *-isMonoid 248 | ; distrib = prop₃ (proj₁ distrib) , prop₃ (proj₂ distrib) 249 | } 250 | } where open CommutativeMonoid 251 | (mkCommutativeMonoid +-commutativeMonoid F _+′_ 0#′) 252 | renaming (isCommutativeMonoid to +-isCommutativeMonoid) 253 | open Monoid (mkMonoid *-monoid F _*′_ 1#′) 254 | renaming (isMonoid to *-isMonoid) 255 | \end{code} 256 | 257 | %<*Semiring> 258 | \begin{code} 259 | module _ (r : Semiring c ℓ) (F : Semiring.Carrier r → Set b) where 260 | open Semiring r hiding (isSemiringWithoutAnnihilatingZero) ; open Dop F 261 | mkSemiring : D₂ _+_ → D₂ _*_ → D₀ 0# → D₀ 1# → Semiring (c ⊔ b) ℓ 262 | mkSemiring _+′_ _*′_ 0#′ 1#′ = record 263 | { isSemiring = record 264 | { isSemiringWithoutAnnihilatingZero = isSemiringWithoutAnnihilatingZero 265 | ; zero = prop₁ zeroˡ , prop₁ zeroʳ 266 | } 267 | } where open SemiringWithoutAnnihilatingZero 268 | (mkSemiringWithoutAnnihilatingZero 269 | semiringWithoutAnnihilatingZero F _+′_ _*′_ 0#′ 1#′) 270 | \end{code} 271 | % 272 | 273 | \begin{code} 274 | module _ (r : CommutativeSemiring c ℓ) (F : CommutativeSemiring.Carrier r → Set b) where 275 | open CommutativeSemiring r hiding (isSemiring) ; open Dop F 276 | mkCommutativeSemiring : D₂ _+_ → D₂ _*_ → D₀ 0# → D₀ 1# 277 | → CommutativeSemiring (c ⊔ b) ℓ 278 | mkCommutativeSemiring _+′_ _*′_ 0#′ 1#′ = record 279 | { isCommutativeSemiring = 280 | record { isSemiring = isSemiring 281 | ; *-comm = prop₂ *-comm 282 | } 283 | } where open Semiring (mkSemiring semiring F _+′_ _*′_ 0#′ 1#′) 284 | 285 | module _ (r : CancellativeCommutativeSemiring c ℓ) 286 | (F : CancellativeCommutativeSemiring.Carrier r → Set b) where 287 | open CancellativeCommutativeSemiring r hiding (isCommutativeSemiring) 288 | open Dop F 289 | mkCancellativeCommutativeSemiring : D₂ _+_ → D₂ _*_ → D₀ 0# → D₀ 1# 290 | → CancellativeCommutativeSemiring (c ⊔ b) ℓ 291 | mkCancellativeCommutativeSemiring _+′_ _*′_ 0#′ 1#′ = record 292 | { isCancellativeCommutativeSemiring = record 293 | { isCommutativeSemiring = isCommutativeSemiring 294 | ; *-cancelˡ-nonZero = λ (y , _) (z , _) q r → *-cancelˡ-nonZero y z q r 295 | } 296 | } where open CommutativeSemiring 297 | (mkCommutativeSemiring commutativeSemiring F _+′_ _*′_ 0#′ 1#′) 298 | 299 | module _ (r : Ring c ℓ) (F : Ring.Carrier r → Set b) where 300 | open Ring r hiding (+-isAbelianGroup; *-isMonoid) ; open Dop F 301 | mkRing : D₂ _+_ → D₂ _*_ → D₁ (-_) → D₀ 0# → D₀ 1# → Ring (c ⊔ b) ℓ 302 | mkRing _+′_ _*′_ -′_ 0#′ 1#′ = record 303 | { isRing = record 304 | { +-isAbelianGroup = +-isAbelianGroup 305 | ; *-isMonoid = *-isMonoid 306 | ; distrib = prop₃ (proj₁ distrib) , prop₃ (proj₂ distrib) 307 | ; zero = prop₁ zeroˡ , prop₁ zeroʳ 308 | } 309 | } where open AbelianGroup (mkAbelianGroup +-abelianGroup F _+′_ 0#′ -′_) 310 | renaming (isAbelianGroup to +-isAbelianGroup) 311 | open Monoid (mkMonoid *-monoid F _*′_ 1#′) 312 | renaming (isMonoid to *-isMonoid) 313 | 314 | module _ (r : CommutativeRing c ℓ) (F : CommutativeRing.Carrier r → Set b) where 315 | open CommutativeRing r hiding (isRing) ; open Dop F 316 | mkCommutativeRing : D₂ _+_ → D₂ _*_ → D₁ (-_) → D₀ 0# → D₀ 1# 317 | → CommutativeRing (c ⊔ b) ℓ 318 | mkCommutativeRing _+′_ _*′_ _-′ 0#′ 1#′ = record 319 | { isCommutativeRing = record { isRing = isRing ; *-comm = prop₂ *-comm } 320 | } where open Ring (mkRing ring F _+′_ _*′_ _-′ 0#′ 1#′) 321 | 322 | \end{code} 323 | -------------------------------------------------------------------------------- /Related work.md: -------------------------------------------------------------------------------- 1 | ## Related work 2 | 3 | * [*Combining predicate transformer semantics for effects: a case study in parsing regular languages*]\: 4 | * ... 5 | 6 | * [*Equational Reasoning about Formal Languages in Coalgebraic Style*]\: 7 | * Focuses on coalgebraic/coinductive techniques. 8 | * Relates to automata theory. 9 | * Implements the same vocabulary on tries as in my paper, but does not prove consistency with a simpler, less operational specification. 10 | * Equivalence is defined via bisimulation, and proofs are coinductive, in terms of tries. 11 | 12 | * [*A Typed, Algebraic Approach to Parsing*]\: 13 | 14 |
15 | In this paper, we recall the definition of the context-free expressions (or μ-regular expressions), an algebraic presentation of the context-free languages. Then, we define a core type system for the context-free expressions which gives a compositional criterion for identifying those context-free expressions which can be parsed unambiguously by predictive algorithms in the style of recursive descent or LL(1). 16 | 17 | Next, we show how these typed grammar expressions can be used to derive a parser combinator library which both guarantees linear-time parsing with no backtracking and single-token lookahead, and which respects the natural denotational semantics of context-free expressions. Finally, we show how to exploit the type information to write a staged version of this library, which produces dramatic increases in performance, even outperforming code generated by the standard parser generator tool ocamlyacc. 18 |19 | 20 | Doesn't seem strongly relevant, since it uses simple types and doesn't capture correctness. 21 | On the other hand, it may have some helpful pointers for extending from regular languages to context-free. 22 | 23 | > This paper describes an implementation of Harper's continuation-based regular-expression matcher as a strong functional program in Cedille; i.e., Cedille statically confirms termination of the program on all inputs. The approach uses neither dependent types nor termination proofs. Instead, a particular interface dubbed a recursion universe is provided by Cedille, and the language ensures that all programs written against this interface terminate. Standard polymorphic typing is all that is needed to check the code against the interface. This answers a challenge posed by Bove, Krauss, and Sozeau. 24 | 25 | * *Strong Functional Pearl: Harper’s Regular-Expression Matcher in Cedille*: 26 | * Abstract: 27 | > This paper describes an implementation of Harper's continuation-based regular-expression matcher as a strong functional program in Cedille; i.e., Cedille statically confirms termination of the program on all inputs. The approach uses neither dependent types nor termination proofs. Instead, a particular interface dubbed a recursion universe is provided by Cedille, and the language ensures that all programs written against this interface terminate. Standard polymorphic typing is all that is needed to check the code against the interface. This answers a challenge posed by Bove, Krauss, and Sozeau. 28 | 29 | * [*Intrinsic Verification of a Regular Expression Matcher*]\: 30 | 31 | * Quote: 32 | > When coding a program/proof in a dependently typed language, there is a choice between "extrinsic" verification (write the simply-typed code, and then prove it correct) and "intrinsic verification" (fuse the program and proof into one function, with a dependent type). 33 | 34 | * Uses Agda 35 | * Relies on defunctionalization to prove termination 36 | * Bob H's original formulation identifies a termination bug. 37 | "The problem can be fixed by restricting the domain of the function to standard regular expressions, which have no Kleene-stared subexpressions that accept the empty string, and then using a preprocessor translation to solve the original problem." 38 | In this paper, they "use a syntactic criterion, generating standard regular expressions by literals, concatenation, alternation, and Kleene plus (one or more occurrences) instead of Kleene star (zero or more occurrences), and omitting the empty string regexp ε." 39 | * The authors define an indexed data type `_∈Ls_` of regular expressions that is quite similar to my few definitions. 40 | * They relate general regular expressions to standard ones. 41 | 42 | * [*Total parser combinators*]\: 43 | * Abstract: 44 |
45 | A monadic parser combinator library which guarantees termination of parsing, while still allowing many forms of left recursion, is described. The library's interface is similar to those of many other parser combinator libraries, with two important differences: one is that the interface clearly specifies which parts of the constructed parsers may be infinite, and which parts have to be finite, using dependent types and a combination of induction and coinduction; and the other is that the parser type is unusually informative. 46 | 47 | The library comes with a formal semantics, using which it is proved that the parser combinators are as expressive as possible. The implementation is supported by a machine-checked correctness proof. 48 |49 | * Has a [GitHub repo](https://github.com/nad/parser-combinators). 50 | * Introduction has many references for parser combinators. 51 | * "In Sections 3.3 and 4.2 Brzozowski derivative operators (Brzozowski 1964) are implemented for recognisers and parsers, and in Sections 3.4 and 4.3 these operators are used to characterise recogniser and parser equivalence coinductively. Rutten (1998) performs similar tasks for regular expressions." 52 | * "Agda is a total language. This means that all computations of inductive type must be terminating, and that all computations of coinductive type must be productive. A computation is productive if the computation of the next constructor is always terminating, so even though an infinite colist cannot be computed in finite time we know that the computation of any finite prefix has to be terminating. For types which are partly inductive and partly coinductive the inductive parts must always be computable in finite time, while the coinductive parts must always be productively computable." 53 | * This paper seems quite relevant and worth studying. 54 | It doesn't appear to be as elegant, but then it is tackling context-free languages. 55 | 56 | * [*Certified Parsing of Regular Languages*]\: 57 | * Abstract: 58 |
59 | We report on a certified parser generator for regular languages using the Agda programming language. Specifically, we programmed a transformation of regular expressions into a Boolean-matrix based representation of nondeterministic finite automata (NFAs). And we proved (in Agda) that a string matches a regular expression if and only if the NFA accepts it. The proof of the if-part is effectively a function turning acceptance of a string into a parse tree while the only-if part gives a function turning rejection into a proof of impossibility of a parse tree. 60 |61 | * Similar language definition as predicate but via an inductive type of regular expressions and another of membership proofs. 62 | * Uses boolean matrices to represent NFA. 63 | "We parse strings with nondeterministic finite automata and represent them in terms of Boolean matrices." 64 | I'd like to understand this technique, and consider replacing the booleans with propositions. 65 | 66 | * [*Regular Expressions in Agda*]\: 67 | * Similar to my work in some ways, but `ν` and `δ` are *defined* on regular expressions instead of proved. 68 | * It's a short paper, but I'm unsure what the authors really proved. 69 | * In the conclusion section: 70 |
71 | However, we note a particular defficiency in our program. That is the inability to prove the non-membership of some string in a given language. Indeed, we prove membership by constructing an element of the dependent data type `x ∈⌣⟦E⟧` in doing this we have to be able to construct an element of the datatype e0 = e/a at each step of the induction, to do a complete pattern matching we need to include the cases when this construction fails (i.e. the `Nothing` case). The fact that we couldn’t construct an element of the datatype doesn’t execulde the possibility that `e′` is a derviative of `e` to a since `e′` might be semanticaly equivalent to `e/a` but syntactically different. 72 |73 | 74 | * [*Certified Context-Free Parsing: A formalisation of Valiant's Algorithm in Agda] 75 | * [*A theory of changes for higher-order languages---incrementalizing λ-calculi by static differentiation*] also used dependent types, though de-emphasized in the paper. 76 | * Derivatives of data types, starting with Brzozowski. 77 | Show it to be a special case. 78 | * AD 79 | * Convolution 80 | * *A play on regular expressions*. 81 | Doesn't seem very relevant. 82 | * [*A constructive theory of regular languages in Coq*]\: 83 |
84 | We present a formal constructive theory of regular languages consisting of about 1400 lines of Coq/Ssreflect. As representations we consider regular expressions, deterministic and nondeterministic automata, and Myhill and Nerode partitions. We construct computable functions translating between these representations and show that equivalence of representations is decidable. We also establish the usual closure properties, give a minimization algorithm for DFAs, and prove that minimal DFAs are unique up to state renaming. Our development profits much from Ssreflect’s support for finite types and graphs. 85 |86 | 87 | * From agda-stdlib/CHANGELOG/v1.2.md: 88 | "A version of the `Reflects` idiom, as seen in SSReflect, has been introduced in `Relation.Nullary`." 89 | Apparently, SSReflect ("Small Scale Reflection") is a Coq extension. 90 | See [*The SSReflect proof language — Coq documentation*](https://coq.inria.fr/refman/proof-engine/ssreflect-proof-language.html) and [*An introduction to small scale reflection in Coq*](https://jfr.unibo.it/article/view/1979). 91 | * [*Formal Languages, Formally and Coinductively*]\: 92 |
93 | Traditionally, formal languages are defined as sets of words. More recently, the alternative coalgebraic or coinductive representation as infinite tries, i.e., prefix trees branching over the alphabet, has been used to obtain compact and elegant proofs of classic results in language theory. In this article, we study this representation in the Isabelle proof assistant. We define regular operations on infinite tries and prove the axioms of Kleene algebra for those operations. Thereby, we exercise corecursion and coinduction and confirm the coinductive view being profitable in formalizations, as it improves over the set-of-words view with respect to proof automation. 94 |95 | 96 | * [*Generic Level Polymorphic N-ary Functions*] 97 | * [*Fun with Semirings*] 98 | * Semiring parsing 99 | 100 | 101 | 102 | [*Equational Reasoning about Formal Languages in Coalgebraic Style*]: http://www.cse.chalmers.se/~abela/jlamp17.pdf "paper by Andreas Abel (2017)" 103 | 104 | [*Formal Languages, Formally and Coinductively*]: https://arxiv.org/abs/1611.09633 "paper by Dmitriy Traytel (2017)" 105 | 106 | [*Proof Pearl: Regular Expression Equivalence and Relation Algebra*]: https://www21.in.tum.de/~krauss/papers/rexp.pdf "paper by Alexander Krauss and Tobias Nipkow (2012)" 107 | 108 | [*Generic Level Polymorphic N-ary Functions*]: https://gallais.github.io/pdf/tyde19.pdf "paper by Guillaume Allais (2019)" 109 | 110 | [*Fun with Semirings: A functional pearl on the abuse of linear algebra*]: http://stedolan.net/research/semirings.pdf "paper by Stephen Dolan (2013)" 111 | 112 | [*Regular Expressions in Agda*]: https://itu.dk/people/basm/report.pdf "paper by Alexandre Agular and Bassel Mannaa (2009)" 113 | 114 | [*Certified Context-Free Parsing: A formalisation of Valiant's Algorithm in Agda]: http://arxiv.org/abs/1601.07724 "paper by Jean-Philippe Bernardy and Patrik Jansson (2016)" 115 | 116 | [*A theory of changes for higher-order languages---incrementalizing λ-calculi by static differentiation*]: http://inc-lc.github.io/resources/pldi14-ilc-author-final.pdf "paper by Yufei Cai, Paolo G. Giarrusso, Tillmann Rendel, and Klaus Ostermann (2014)" 117 | 118 | [*Certified Parsing of Regular Languages*]: https://citeseerx.ist.psu.edu/viewdoc/summary?doi=10.1.1.571.724 "paper by Denis Firsov , Tarmo Uustalu" 119 | 120 | [*Intrinsic Verification of a Regular Expression Matcher*]: http://cattheory.com/regexp2016.pdf "Joomy Korkut, Maksim Trifunovski, and Daniel R. Licata (2016)" 121 | 122 | [*A Typed, Algebraic Approach to Parsing*]: https://www.cl.cam.ac.uk/~nk480/parsing.pdf "paper by Neel Krishnaswami and Jeremy Yallop (2019)" 123 | 124 | [*Total parser combinators*]: http://www.cse.chalmers.se/~nad/publications/danielsson-parser-combinators.html "paper by Nils Anders Danielsson (2015)" 125 | 126 | [*A constructive theory of regular languages in Coq*]: https://www.ps.uni-saarland.de/~doczkal/regular/ConstructiveRegularLanguages.pdf "paper by Christian Doczkal, Jan-Oliver Kaiser, and Gert Smolka (2013)" 127 | 128 | [*Combining predicate transformer semantics for effects: a case study in parsing regular languages*]: https://arxiv.org/abs/2005.00197 "paper by Anne Baanen Wouter Swierstra (2020)" 129 | -------------------------------------------------------------------------------- /Inverses.lagda: -------------------------------------------------------------------------------- 1 | \sectionl{Some isomorphism properties} 2 | 3 | \note{Not to be included in the paper other than for testing.} 4 | 5 | \begin{code} 6 | 7 | {-# OPTIONS --safe #-} 8 | -- {-# OPTIONS --without-K #-} -- See f⁻¹∘f∘f⁻¹ below. 9 | 10 | open import Level 11 | 12 | module Inverses {ℓ : Level} where 13 | 14 | open import Data.Unit using () renaming (tt to tt₀) 15 | open import Algebra.Core 16 | open import Data.Product 17 | open import Data.Product.Properties 18 | open import Data.Product.Algebra 19 | open import Data.Sum hiding (map₂) 20 | open import Data.Sum.Algebra 21 | open import Data.List 22 | open import Data.List.Properties 23 | open import Function using (id; _∘_; const; flip; _↔_; Inverse; mk↔′; mk↔) 24 | open import Relation.Binary using (Rel; IsEquivalence) 25 | open import Relation.Binary.PropositionalEquality 26 | open import Function.Equivalence using (_⇔_; equivalence) 27 | 28 | open import Function.Properties.Inverse using (↔-isEquivalence) 29 | module ↔Eq {ℓ} = IsEquivalence {ℓ = ℓ} ↔-isEquivalence 30 | import Relation.Binary.Reasoning.Setoid as SetoidR 31 | module ↔R {ℓ} = SetoidR {s₂ = ℓ} (record {isEquivalence = ↔-isEquivalence}) 32 | 33 | open import Misc {ℓ} 34 | 35 | open import Predicate {ℓ} 36 | 37 | private 38 | variable 39 | A B C D : Set ℓ 40 | P Q R S : Pred A 41 | 42 | ---- Basics 43 | 44 | -- Function-lifted _↔_ 45 | \end{code} 46 | %<*ext-iso> 47 | \AgdaTarget{\_⟷\_, ⟷} 48 | \begin{code}[hide] 49 | infix 3 _⟷_ 50 | _⟷_ : Rel (Pred A) ℓ 51 | \end{code} 52 | \begin{code} 53 | P ⟷ Q = ∀ {w} → P w ↔ Q w 54 | \end{code} 55 | % 56 | 57 | \begin{code} 58 | infix 3 _⟺_ 59 | _⟺_ : Rel (Pred A) ℓ 60 | P ⟺ Q = ∀ {w} → P w ⇔ Q w 61 | \end{code} 62 | 63 | \begin{code} 64 | 65 | ⟷-isEquivalence : IsEquivalence (_⟷_ {A = A}) 66 | ⟷-isEquivalence = record 67 | { refl = ↔Eq.refl 68 | ; sym = λ f⟷g → ↔Eq.sym f⟷g 69 | ; trans = λ f⟷g g⟷k → ↔Eq.trans f⟷g g⟷k 70 | } 71 | 72 | module ⟷Eq {A : Set ℓ} = IsEquivalence (⟷-isEquivalence {A = A}) 73 | import Relation.Binary.Reasoning.Setoid as SetoidR 74 | module ⟷R {A : Set ℓ} = SetoidR (record {isEquivalence = ⟷-isEquivalence {A = A}}) 75 | \end{code} 76 | 77 | \begin{code} 78 | ≡↔ : A ≡ B → A ↔ B 79 | ≡↔ refl = ↔Eq.refl 80 | 81 | sym↔ : ∀ {a b : A} → a ≡ b ↔ b ≡ a 82 | sym↔ = mk↔′ sym sym (λ {refl → refl}) (λ {refl → refl}) 83 | 84 | -- Maybe give this one a more specific name and add ?? : a ≡ b → b ≡ c ↔ a ≡ c 85 | trans↔ : ∀ {a b c : A} → b ≡ c → a ≡ b ↔ a ≡ c 86 | trans↔ refl = 87 | mk↔′ (λ {refl → refl}) (λ {refl → refl}) (λ {refl → refl}) (λ {refl → refl}) 88 | \end{code} 89 | 90 | \subsection{Products and sums} 91 | 92 | %<*xCongRSig> 93 | \begin{code} 94 | ×-congʳ : A ↔ B → (A × C) ↔ (B × C) 95 | \end{code} 96 | % 97 | \begin{code} 98 | ×-congʳ A↔B = ×-cong A↔B ↔Eq.refl 99 | 100 | ×-congˡ : C ↔ D → (A × C) ↔ (A × D) 101 | ×-congˡ C↔D = ×-cong ↔Eq.refl C↔D 102 | \end{code} 103 | 104 | %<*uCongRSig> 105 | \begin{code} 106 | ⊎-congʳ : A ↔ B → (A ⊎ C) ↔ (B ⊎ C) 107 | \end{code} 108 | % 109 | \begin{code} 110 | ⊎-congʳ A↔B = ⊎-cong A↔B ↔Eq.refl 111 | 112 | ⊎-congˡ : C ↔ D → (A ⊎ C) ↔ (A ⊎ D) 113 | ⊎-congˡ C↔D = ⊎-cong ↔Eq.refl C↔D 114 | 115 | ∃-distrib-⟨×⟩ : ∃ (P ⟨×⟩ Q) ↔ (∃ P × ∃ Q) 116 | ∃-distrib-⟨×⟩ {P = P}{Q = Q} = mk↔′ 117 | (λ ((a , b) , (Pa , Qb)) → ((a , Pa) , (b , Qb))) 118 | (λ ((a , Pa) , (b , Qb)) → ((a , b) , (Pa , Qb))) 119 | (λ _ → refl) 120 | (λ _ → refl) 121 | 122 | ×-transpose : ((A × B) × (C × D)) ↔ ((A × C) × (B × D)) 123 | ×-transpose = ∃-distrib-⟨×⟩ 124 | 125 | ∃-distrib-∪ : ∃ (P ∪ Q) ↔ (∃ P ⊎ ∃ Q) 126 | ∃-distrib-∪ {P = P}{Q} = mk↔′ 127 | (λ {(a , inj₁ p) → inj₁ (a , p) ; (a , inj₂ q) → inj₂ (a , q)}) 128 | (λ {(inj₁ (a , p)) → a , inj₁ p ; (inj₂ (a , q)) → a , inj₂ q}) 129 | (λ {(inj₁ (a , p)) → refl ; (inj₂ (a , q)) → refl}) 130 | (λ {(a , inj₁ p) → refl ; (a , inj₂ q) → refl}) 131 | 132 | ∃-distrib-∅ : Σ A (λ _ → ⊥) ↔ ⊥ 133 | ∃-distrib-∅ = mk↔′ (λ ()) (λ ()) (λ ()) (λ ()) 134 | 135 | ∃-distrib-⊤ : Σ A (λ _ → ⊤) ↔ A 136 | ∃-distrib-⊤ = ×-identityʳ _ _ 137 | 138 | ∃-distrib-∩ : ∃ (λ (a , b) → P a × Q b) ↔ (∃ P × ∃ Q) 139 | ∃-distrib-∩ = mk↔′ 140 | (λ ((a , b) , (Pa , Qb)) → (a , Pa) , (b , Qb)) 141 | (λ ((a , Pa) , (b , Qb)) → (a , b) , (Pa , Qb)) 142 | (λ ((a , Pa) , (b , Qb)) → refl) 143 | (λ ((a , b) , (Pa , Qb)) → refl) 144 | 145 | ∃-×-constˡ : ∃ (λ z → A × Q z) ↔ (A × ∃ Q) 146 | ∃-×-constˡ = mk↔′ 147 | (λ (z , (a , Qz)) → a , (z , Qz)) 148 | (λ (a , (z , Qz)) → z , (a , Qz)) 149 | (λ (a , (z , Qz)) → refl) 150 | (λ (z , (a , Qz)) → refl) 151 | 152 | ∃-×-constʳ : ∃ (λ z → P z × B) ↔ (∃ P × B) 153 | ∃-×-constʳ = mk↔′ 154 | (λ (z , (Pz , b)) → (z , Pz) , b) 155 | (λ ((z , Pz) , b) → z , (Pz , b)) 156 | (λ ((z , Pz) , b) → refl) 157 | (λ (z , (Pz , b)) → refl) 158 | 159 | ×-distrib-∃ˡ : (A × ∃ Q) ↔ ∃ (λ v → A × Q v) 160 | ×-distrib-∃ˡ = mk↔′ 161 | (λ (a , b , Qb) → b , a , Qb) 162 | (λ (b , a , Qb) → a , b , Qb) 163 | (λ (b , a , Qb) → refl) 164 | (λ (a , b , Qb) → refl) 165 | 166 | ×-distrib-∃ʳ : (∃ P × B) ↔ ∃ (λ u → P u × B) 167 | ×-distrib-∃ʳ = mk↔′ 168 | (λ ((a , Pa) , b) → a , Pa , b) 169 | (λ ( a , Pa , b) → (a , Pa) , b) 170 | (λ ( a , Pa , b) → refl) 171 | (λ ((a , Pa) , b) → refl) 172 | 173 | \end{code} 174 | 175 | \subsection{Existential quantification} 176 | %<*exCongSig> 177 | \begin{code} 178 | ∃-cong : (∀ {w} → P w ↔ Q w) → (∃ P ↔ ∃ Q) 179 | \end{code} 180 | % 181 | \begin{code} 182 | ∃-cong P↔Q = mk↔′ (map₂ J.f) (map₂ J.f⁻¹) 183 | (λ (a , d) → Inverse.f Σ-≡,≡↔≡ (refl , J.inverseˡ d) ) 184 | (λ (a , c) → Inverse.f Σ-≡,≡↔≡ (refl , J.inverseʳ c) ) 185 | where 186 | module J {t} = Inverse (P↔Q {t} ) 187 | \end{code} 188 | 189 | \begin{code} 190 | ∃₂-cong : ∀ {f g : A → B → Set ℓ} → (∀ {a b} → f a b ↔ g a b) → ∃₂ f ↔ ∃₂ g 191 | ∃₂-cong f↔g = ∃-cong (∃-cong f↔g) 192 | 193 | ∃₃ : ∀ {A : Set ℓ} {B : A → Set ℓ} {C : (x : A) → B x → Set ℓ} 194 | (D : (x : A) (y : B x) → C x y → Set ℓ) → Set ℓ 195 | ∃₃ D = ∃ λ a → ∃₂ λ b c → D a b c 196 | 197 | ∃₃-cong : ∀ {D E : A → B → C → Set ℓ} → (∀ {a b c} → D a b c ↔ E a b c) → ∃₃ D ↔ ∃₃ E 198 | ∃₃-cong D↔E = ∃-cong (∃₂-cong D↔E) 199 | \end{code} 200 | 201 | Existentials with various domain shapes: 202 | \begin{code} 203 | ∃⊥ : ∀ {P : ⊥ → Set ℓ} → ∃ P ↔ ⊥ 204 | ∃⊥ = mk↔′ (λ ()) (λ ()) (λ ()) (λ ()) 205 | 206 | ∃⊎ : ∀ {P : A ⊎ B → Set ℓ} → ∃ P ↔ (∃ (P ∘ inj₁) ⊎ ∃ (P ∘ inj₂)) 207 | ∃⊎ = mk↔′ 208 | (λ { (inj₁ b , z) → inj₁ (b , z) ; (inj₂ c , z) → inj₂ (c , z) }) 209 | (λ { (inj₁ (b , z)) → inj₁ b , z ; (inj₂ (c , z)) → inj₂ c , z }) 210 | (λ { (inj₁ _) → refl ; (inj₂ _) → refl }) 211 | (λ { (inj₁ _ , _) → refl ; (inj₂ _ , _) → refl }) 212 | 213 | ∃⊤ : ∀ {P : ⊤ → Set ℓ} → ∃ P ↔ P tt 214 | ∃⊤ = mk↔′ 215 | (λ { (tt , Ptt) → Ptt }) 216 | (λ { Ptt → tt , Ptt }) 217 | (λ { Ptt → refl }) 218 | (λ { (tt , Ptt) → refl }) 219 | 220 | open import Data.Bool 221 | 222 | ∃Bool : ∀ {P : Bool → Set ℓ} → ∃ P ↔ (P false ⊎ P true) 223 | ∃Bool = mk↔′ 224 | (λ { (false , Pfalse) → inj₁ Pfalse ; (true , Ptrue) → inj₂ Ptrue }) 225 | (λ { (inj₁ Pfalse) → false , Pfalse ; (inj₂ Ptrue) → true , Ptrue }) 226 | (λ { (inj₁ Pfalse) → refl ; (inj₂ Ptrue) → refl }) 227 | (λ { (false , Pfalse) → refl ; (true , Ptrue) → refl }) 228 | 229 | \end{code} 230 | 231 | \begin{code} 232 | ∃sub : ∀ {f : A → B} {P : B → Set ℓ} 233 | → (∃ λ b → P (f b)) ↔ (∃ λ c → P c × ∃ λ b → c ≡ f b) 234 | ∃sub {f = f}{P} = mk↔′ 235 | (λ { (b , Pc) → f b , Pc , b , refl }) 236 | (λ { (.(f b) , Pc , b , refl) → b , Pc }) 237 | (λ { (.(f b) , Pc , b , refl) → refl }) 238 | (λ { (b , Pc) → refl }) 239 | 240 | ∃×↔∃∃ : ∀ {P : A × B → Set ℓ} → ∃ P ↔ ∃₂ (curry P) 241 | ∃×↔∃∃ = mk↔′ 242 | (λ {((b , c) , Pb,c) → b , c , Pb,c}) 243 | (λ {(b , c , Pb,c) → (b , c) , Pb,c}) 244 | (λ {(b , c , Pb,c) → refl}) 245 | (λ {((b , c) , Pb,c) → refl}) 246 | 247 | ∃× : ∃⇃ (λ (a , b) → P a × Q b) ↔ (∃ P × ∃ Q) 248 | ∃× = mk↔′ 249 | (λ {((a , b) , (fa , gb)) → (a , fa) , (b , gb)}) 250 | (λ {((a , fa) , (b , gb)) → (a , b) , (fa , gb)}) 251 | (λ _ → refl) 252 | (λ _ → refl) 253 | 254 | ∃≡proj₁ : ∀ {ℓ₁ ℓ₂}{A : Set ℓ₁} {B : A → Set ℓ₂} {a : A} 255 | → ∃ (λ (ab : Σ A B) → a ≡ proj₁ ab) ↔ B a 256 | ∃≡proj₁ {a = a} = mk↔′ 257 | (λ { ((.a , b) , refl) → b }) 258 | (λ { b → (a , b) , refl }) 259 | (λ { b → refl }) 260 | (λ { ((.a , b) , refl) → refl }) 261 | 262 | ∃≡ : ∀ {x : A} → ∃ (x ≡_) ↔ ⊤ 263 | ∃≡ {x = x} = mk↔′ (λ {(.x , refl) → tt}) (λ {tt → x , refl}) 264 | (λ {tt → refl}) (λ {(.x , refl) → refl}) 265 | \end{code} 266 | 267 | %<*exElimSig> 268 | \begin{code} 269 | ∃≡ˡ : ∀ {z} → (∃ λ u → u ≡ z × P u) ↔ P z 270 | \end{code} 271 | % 272 | \begin{code} 273 | ∃≡ˡ {z = z} = mk↔′ (λ {(.z , refl , Pz) → Pz} ) (λ Pz → z , refl , Pz) 274 | (λ _ → refl) (λ {(.z , refl , Pz) → refl}) 275 | 276 | -- Where should ∃¹ and ∃² go? 277 | 278 | ∃¹ : (A → Pred B) → Pred B 279 | ∃¹ f y = ∃ λ x → f x y 280 | 281 | ∃² : (A → B → Pred C) → Pred C 282 | ∃² f z = ∃₂ λ x y → f x y z 283 | 284 | ∃¹-cong : ∀ {f g : A → Pred B} → (∀ {a} → f a ⟷ g a) → ∃¹ f ⟷ ∃¹ g 285 | ∃¹-cong f⟷g = ∃-cong f⟷g 286 | 287 | -- ∃¹-cong {D = D}{E} D⟷E {z} = 288 | -- begin 289 | -- ∃¹ D z 290 | -- ≡⟨⟩ 291 | -- (∃ λ x → D x z) 292 | -- ≈⟨ ∃-cong D⟷E ⟩ 293 | -- (∃ λ x → E x z) 294 | -- ≡⟨⟩ 295 | -- ∃¹ E z 296 | -- ∎ 297 | -- where open ↔R 298 | 299 | -- ∃≡ʳ′ : ∀ {A : Set ℓ} {P : A → Set ℓ} {z} → (∃ λ u → P u × u ≡ z) ↔ P z 300 | -- ∃≡ʳ′ {P = P} {z = z} = 301 | -- begin 302 | -- (∃ λ u → P u × u ≡ z) 303 | -- ≈⟨ ∃-cong (×-comm _ _) ⟩ 304 | -- (∃ λ u → u ≡ z × P u) 305 | -- ≈⟨ ∃≡ˡ ⟩ 306 | -- P z 307 | -- ∎ 308 | -- where open ↔R 309 | 310 | -- I'd rather let A and B and the result of P have different sorts, but I got a 311 | -- type error. If and when motivated, try again. Meanwhile, the following 312 | -- version is level-polymorphic. 313 | 314 | ∃≡ʳ : ∀ {P : A → Set ℓ} {z} → (∃ λ u → P u × u ≡ z) ↔ P z 315 | ∃≡ʳ {z = z} = mk↔′ 316 | (λ {(u , Pz , refl) → Pz}) 317 | (λ Pz → z , Pz , refl) 318 | (λ _ → refl) 319 | (λ {(u , Pz , refl) → refl}) 320 | 321 | -- -- I think the level limitations are coming from the IsEquivalence fields 322 | -- -- limiting levels in sym and trans. Can we do equational-style reasoning 323 | -- -- without these limitations? 324 | 325 | -- ∃₂≡ˡ′ : ∀ {A : Set ℓ} {B : Set ℓ} {P : A → B → Set ℓ} {y z} 326 | -- → (∃₂ λ u v → (u ≡ y × v ≡ z) × P u v) ↔ P y z 327 | -- ∃₂≡ˡ′ {a = a} {b} {P = P} {y = y} {z} = 328 | -- begin 329 | -- (∃₂ λ u v → (u ≡ y × v ≡ z) × P u v) 330 | -- ≈⟨ ∃₂-cong (×-assoc _ _ _ _) ⟩ 331 | -- (∃₂ λ u v → u ≡ y × (v ≡ z × P u v)) 332 | -- ≈⟨ ∃∃↔∃∃ _ ⟩ 333 | -- (∃₂ λ v u → u ≡ y × (v ≡ z × P u v)) 334 | -- ≈⟨ ∃-cong ∃≡ˡ ⟩ 335 | -- (∃ λ v → v ≡ z × P y v) 336 | -- ≈⟨ ∃≡ˡ ⟩ 337 | -- P y z 338 | -- ∎ 339 | -- where open ↔R 340 | 341 | -- -- I'd rather let A and B and the result of P have different sorts, but I got a 342 | -- -- type error. If and when motivated, try again. Meanwhile, the following 343 | -- -- version is level-polymorphic. 344 | 345 | -- ∃₂≡ˡ : ∀ {A : Set ℓ} {B : Set ℓ} {P : A → B → Set c} {y z} 346 | -- → (∃₂ λ u v → (u ≡ y × v ≡ z) × P u v) ↔ P y z 347 | -- ∃₂≡ˡ {y = y} {z} = mk↔′ 348 | -- (λ {(.y , .z , (refl , refl) , Pyz) → Pyz}) 349 | -- (λ Pyz → y , z , (refl , refl) , Pyz) 350 | -- (λ Pyz → refl) 351 | -- (λ {(.y , .z , (refl , refl) , Pyz) → refl}) 352 | 353 | -- ∃₂≡ʳ : ∀ {A : Set ℓ} {B : Set ℓ} {P : A → B → Set c} {y z} 354 | -- → (∃₂ λ u v → P u v × (u ≡ y × v ≡ z)) ↔ P y z 355 | -- ∃₂≡ʳ {y = y} {z} = mk↔′ 356 | -- (λ {(.y , .z , Pyz , (refl , refl)) → Pyz}) 357 | -- (λ Pyz → y , z , Pyz , (refl , refl)) 358 | -- (λ Pyz → refl) 359 | -- (λ {(.y , .z , Pyz , (refl , refl)) → refl}) 360 | 361 | ∃≡² : ∀ {P Q : A → Set ℓ} {v} → (∃ λ p → P p × p ≡ v × Q p) ↔ (P v × Q v) 362 | ∃≡² {v = v} = mk↔′ 363 | (λ {(p , Pp , refl , Qp) → Pp , Qp}) 364 | (λ {(Pv , Qv) → v , Pv , refl , Qv}) 365 | (λ {(Pv , Qv) → refl}) 366 | (λ {(p , Pp , refl , Qp) → refl}) 367 | 368 | -- distrib-∃ˡ : ∀ {A : Set ℓ}{G : Set d → Set ℓ} → (A × ∃ G) ↔ ∃ (λ u → A × G u) -- ∃ (A *ₗ G) 369 | -- distrib-∃ˡ {A = A}{G = G} = mk↔′ 370 | -- (λ (a , u , b) → u , a , b) 371 | -- (λ (u , a , b) → a , u , b) 372 | -- (λ _ → refl) 373 | -- (λ _ → refl) 374 | 375 | -- -- TODO: Maybe redefine distrib-∃ˡ via ∃∃↔∃∃ 376 | 377 | \end{code} 378 | 379 | \subsection{Predicates} 380 | 381 | \begin{code} 382 | 383 | ⟨×⟩-cong : ∀ {P R : Pred A} {Q S : Pred B} → P ⟷ R → Q ⟷ S → (P ⟨×⟩ Q) ⟷ (R ⟨×⟩ S) 384 | -- ⟨×⟩-cong P⟷R Q⟷S = ×-cong P⟷R Q⟷S 385 | ⟨×⟩-cong P⟷R Q⟷S {a , b} = ×-cong (P⟷R {a}) (Q⟷S {b}) 386 | 387 | -- ⟨×⟩-cong {P = P}{R}{Q}{S} P⟷R Q⟷S {(a , b)} = let open ↔R in 388 | -- begin 389 | -- (P ⟨×⟩ Q) (a , b) 390 | -- ≡⟨⟩ 391 | -- (P a × Q b) 392 | -- ≈⟨ ×-cong P⟷R Q⟷S ⟩ 393 | -- (R a × S b) 394 | -- ≡⟨⟩ 395 | -- (R ⟨×⟩ S) (a , b) 396 | -- ∎ 397 | 398 | infixr 2 _⟪×⟫_ 399 | _⟪×⟫_ : (A → Pred C) → (B → Pred D) → ((A × B) → Pred (C × D)) 400 | (f ⟪×⟫ g) (a , b) = f a ⟨×⟩ g b 401 | 402 | ∃¹⟪×⟫ : ∀ {f : A → Pred C}{g : B → Pred D} → ∃¹ (f ⟪×⟫ g) ⟷ (∃¹ f ⟨×⟩ ∃¹ g) 403 | ∃¹⟪×⟫ = ∃-distrib-⟨×⟩ 404 | \end{code} 405 | 406 | \begin{code} 407 | module _ {P : A → Set ℓ} (A↔B : A ↔ B) where 408 | open Inverse A↔B 409 | open ≡-Reasoning 410 | private 411 | 412 | g : ∃ P → ∃ (P ∘ f⁻¹) 413 | g (a , Pa) = f a , subst P (sym (inverseʳ a)) Pa 414 | -- g (a , Pa) = f a , {!!} 415 | 416 | g⁻¹ : ∃ (P ∘ f⁻¹) → ∃ P 417 | g⁻¹ (b , Pf⁻¹b) = f⁻¹ b , Pf⁻¹b 418 | 419 | g⁻¹∘g : g⁻¹ ∘ g ≗ id 420 | g⁻¹∘g (a , Pa) = Inverse.f Σ-≡,≡↔≡ 421 | (inverseʳ a , subst-subst-sym (inverseʳ a)) 422 | 423 | -- Two proofs of f⁻¹ (f (f⁻¹ b)) ≡ f⁻¹ b. Uses K. 424 | open import Axiom.UniquenessOfIdentityProofs.WithK 425 | f⁻¹∘f∘f⁻¹ : ∀ b → cong f⁻¹ (inverseˡ b) ≡ inverseʳ (f⁻¹ b) 426 | f⁻¹∘f∘f⁻¹ _ = uip _ _ 427 | 428 | g∘g⁻¹ : g ∘ g⁻¹ ≗ id 429 | g∘g⁻¹ (b , Pf⁻¹b) = 430 | Inverse.f Σ-≡,≡↔≡ 431 | ( inverseˡ b 432 | , (begin 433 | subst (P ∘ f⁻¹) (inverseˡ b) 434 | (subst P (sym (inverseʳ (f⁻¹ b))) Pf⁻¹b) 435 | ≡⟨ subst-∘ (inverseˡ b) ⟩ 436 | subst P (cong f⁻¹ (inverseˡ b)) 437 | (subst P (sym (inverseʳ (f⁻¹ b))) Pf⁻¹b) 438 | ≡⟨ cong (λ z → subst P z _) (f⁻¹∘f∘f⁻¹ b) ⟩ 439 | subst P (inverseʳ (f⁻¹ b)) 440 | (subst P (sym (inverseʳ (f⁻¹ b))) Pf⁻¹b) 441 | ≡⟨ subst-subst-sym (inverseʳ (f⁻¹ b)) ⟩ 442 | Pf⁻¹b 443 | ∎) 444 | ) 445 | 446 | -- Phew! Is there an easier proof? And maybe one that doesn't assume K? 447 | -- Maybe define specialized versions instead. See ∃✶ below. 448 | 449 | ∃dom : ∃ P ↔ ∃ (P ∘ f⁻¹) 450 | ∃dom = mk↔′ g g⁻¹ g∘g⁻¹ g⁻¹∘g 451 | 452 | -- Lists 453 | module _ {X : Set ℓ} where 454 | 455 | private 456 | variable 457 | x y : X 458 | u v w : X ✶ 459 | 460 | ∷≡[]↔ : x ∷ w ≡ [] ↔ ⊥ 461 | ∷≡[]↔ = mk↔′ (λ ()) (λ ()) (λ ()) (λ ()) 462 | 463 | []≡∷↔ : [] ≡ x ∷ w ↔ ⊥ 464 | []≡∷↔ = ↔Eq.trans sym↔ ∷≡[]↔ 465 | 466 | ∷≡∷↔ : y ∷ u ≡ x ∷ w ↔ (y ≡ x × u ≡ w) 467 | ∷≡∷↔ = mk↔′ (λ {refl → refl , refl} ) (λ {(refl , refl) → refl}) 468 | (λ {(refl , refl) → refl}) (λ {refl → refl}) 469 | 470 | \end{code} 471 | %<*emptyAppend> 472 | \begin{code} 473 | []≡++ : ([] ≡ u ++ v) ↔ (u , v) ≡ ([] , []) 474 | \end{code} 475 | % 476 | \begin{code}[hide] 477 | []≡++ {u = []} {[]} = 478 | mk↔′ (λ {refl → refl}) (λ {refl → refl}) 479 | (λ {refl → refl}) (λ {refl → refl}) 480 | []≡++ {u = []} {_ ∷ _} = mk↔′ (λ ()) (λ ()) (λ ()) (λ ()) 481 | []≡++ {u = _ ∷ _} {_} = mk↔′ (λ ()) (λ ()) (λ ()) (λ ()) 482 | 483 | module _ where 484 | open import Closed.Instances using (module Types) 485 | open Types {ℓ} 486 | iso = ✶-star X 487 | open Inverse iso 488 | open ↔R 489 | 490 | ∃✶A : ∀ {P} → ∃ P ↔ ∃ (P ∘ f⁻¹) 491 | ∃✶A = ∃dom iso 492 | 493 | -- ∃✶ : ∀ {P} → ∃ P ↔ (P [] ⊎ ∃⇃ (λ (x , xs) → P (x ∷ xs))) 494 | -- ∃✶ {P = P} = 495 | -- begin 496 | -- ∃ P 497 | -- ≈⟨ ∃dom {A↔B = iso} ⟩ 498 | -- ∃ (P ∘ f⁻¹) 499 | -- ≡⟨⟩ 500 | -- Σ (⊤ ⊎ X × X ✶) 501 | -- (λ z → P ((λ { (inj₁ _) → [] ; (inj₂ (x , xs)) → x ∷ xs }) z)) 502 | -- ≈⟨ {!!} ⟩ 503 | -- (P [] ⊎ ∃⇃ (λ (x , xs) → P (x ∷ xs))) 504 | -- ∎ 505 | 506 | ∃✶ : ∀ {X} {P : X ✶ → Set ℓ} → ∃ P ↔ (P [] ⊎ ∃⇃ λ (x , xs) → P (x ∷ xs)) 507 | ∃✶ {P = P} = mk↔′ 508 | (λ { ([] , P[]) → inj₁ P[] ; (x ∷ xs , Px∷xs) → inj₂ ((x , xs) , Px∷xs) }) 509 | (λ { (inj₁ P[]) → [] , P[] ; (inj₂ ((x , xs) , Px∷xs)) → (x ∷ xs) , Px∷xs }) 510 | (λ { (inj₁ P[]) → refl ; (inj₂ ((x , xs) , Px∷xs)) → refl }) 511 | (λ { ([] , P[]) → refl ; (x ∷ xs , Px∷xs) → refl }) 512 | 513 | \end{code} 514 | 515 | \begin{code} 516 | ---- Equivalence and decidability 517 | 518 | module _ where 519 | 520 | open import Relation.Nullary using (Dec) 521 | open import Relation.Nullary.Decidable using () renaming (map to map?⇔) 522 | open import Relation.Unary using (Decidable) 523 | 524 | ↔→⇔ : (A ↔ B) → (A ⇔ B) 525 | ↔→⇔ A↔B = equivalence f f⁻¹ where open Inverse A↔B 526 | 527 | \end{code} 528 | %<*mapDec> 529 | \begin{code} 530 | map?↔ : A ↔ B → Dec A → Dec B 531 | \end{code} 532 | % 533 | \begin{code} 534 | map?↔ = map?⇔ ∘ ↔→⇔ 535 | \end{code} 536 | %<*mapDecInv> 537 | \AgdaTarget{map?⁻¹} 538 | \begin{code} 539 | map?⁻¹ : B ↔ A → Dec A → Dec B 540 | \end{code} 541 | % 542 | \begin{code} 543 | map?⁻¹ = map?↔ ∘ ↔Eq.sym 544 | 545 | -- map?⁻¹∘ : {X : Set ℓ} → B ↔ A → (X → Dec A) → (X → Dec B) 546 | map?⁻¹∘ : {W : Set ℓ}{A B : W → Set ℓ} 547 | → (B ⟷ A) → ((s : W) → Dec (A s)) → ((s : W) → Dec (B s)) 548 | map?⁻¹∘ iso f x = map?⁻¹ iso (f x) 549 | 550 | \end{code} 551 | -------------------------------------------------------------------------------- /reviews.md: -------------------------------------------------------------------------------- 1 | ICFP 2021 Paper #87 Reviews and Comments 2 | =========================================================================== 3 | 4 | Paper #87 Symbolic and Automatic Differentiation of Languages 5 | 6 | 7 | Review #87A 8 | =========================================================================== 9 | 10 | Overall merit 11 | ------------- 12 | 13 | B. I support accepting this paper but will not champion it. 14 | 15 | Reviewer expertise 16 | ------------------ 17 | 18 | Z. Some familiarity 19 | 20 | Paper summary 21 | ------------- 22 | 23 | This paper fills in a possibly missing link between the list-based semantics of regular expressions and the trie-based coinductive semantics. The work was mechanized in Agda, using sized types to overcome one technical difficulty in defining the concatenation operator for the coinductive semantics. 24 | 25 | Comments for author 26 | ------------------- 27 | 28 | The two semantics used in the paper seem intuitive and I was a bit surprised that the equivalence has not been mechanized before. If the linkage is indeed missing from the literature, then this is a wonderful piece of work to be included. 29 | 30 | I seemed to be able to infer the intended meaning of all the code, and the discussions look reasonable. However, I am a proficient Agda user and some notes on the syntax could help more general audience. 31 | 32 | My major concern is novelty, but unfortunately I am not familiar with the prior work. Given that I was a bit surprised by the absence of prior work, and that there were already so many papers on related topics, my rating is B. I might advocate for this paper more enthusiastically once novelty is confirmed. 33 | 34 | ##### Response to other reviews about running Agda 35 | 36 | I can check all the supplied Agda code, using the latest development versions (the master branches on GitHub) of Agda and the Agda standard library. The entire codebase contains two use of the `TERMINATING` pragma in `Automatic.lagda` which turns off the termination checker for the operations _not_ using sized types (as discussed in L420-432), but otherwise I did not find anything suspicious. As a bonus point, the authors have tried to impose `--without-K` that effectively turns off UIP when possible. I recognize that a note on the versions of Agda and its standard library would help the reader. 37 | 38 | 39 | 40 | Review #87B 41 | =========================================================================== 42 | 43 | Overall merit 44 | ------------- 45 | 46 | B. I support accepting this paper but will not champion it. 47 | 48 | Reviewer expertise 49 | ------------------ 50 | 51 | Z. Some familiarity 52 | 53 | Paper summary 54 | ------------- 55 | 56 | This paper studies type-theoretic accounts of parsing. It generalizes Brzozowsik's derivatives for regular expressions to arbitrary languages and points out that we can recover the original derivatives via reflections. It also finds that the reflection applied to the dual (i.e., coinductive) definition leads to tries. The theory is implemented and proved in Agda. 57 | 58 | I generally like this paper. A type-theoretic view of parsing is an interesting research topic, and it reports several findings. In particular, I was excited to use automatic differentiation for calculating Brzozowski derivatives. However, its presentation is not good. It is like the authors' own memo. It enumerates necessary formulas without providing sufficient explanations. I suspect that the current draft is accessible only by limited experts. 59 | 60 | Comments for author 61 | ------------------- 62 | 63 | * Not all ICFP audiences are familiar with the dependent type theory and Agda (even though most of them have heard them). Recalling basic notations and concepts is very useful. 64 | 65 | * It seems that in Fig. 1 "a language over an alphabet A" is not A* but A* -> Set l. This is at least nonstandard --- a standard definition should be A*. Explain the reason. By the way, can you use Prop instead of Set l? I suspect the former is more intuitive, especially for those who are not experts on type theory. 66 | 67 | * In Fig. 1, while I believe c :: A, I could not figure out the type of scalar, s. 68 | 69 | * This paper carefully distinguished three kinds of equalities, namely propositional equality, isomorphism, and extensional isomorphism. However, I did not understand why that equality should be used for each occurrence. For example, why in Fig. 1 is propositional equality necessary? 70 | 71 | * Section 2 uses A* -> B instead of A* -> Set l. Perhaps, the reason is that all the definitions can be polymorphic on the result, B. However, this difference at least makes the meaning difficult to understand. 72 | 73 | * p5. "isomorphic types or predicates are equivalently decidable": 74 | I could not understand the reason. 75 | 76 | * Figs. 4 and 5 are nearly identical to Figs. 6 and 7. While I could understand that it is because of the duality, I still wonder what is the essential differences. Usually, an inductive definition gives a datatype, whereas a coinductive one does a function (or stream). However, the automatic differentiation requires the concrete construction process of the function even for the coinductive case. Another possible difference is the potential of the infiniteness of the coinductive definition. However, for this particular case, I am not aware of the usefulness of infiniteness. 77 | 78 | * On page 9. it is argued that a difference is a work-sharing between the function and its derivative. I agree with this point but do not think it to be critical. Existing parsing methods based on the derivatives do such work sharing. Moreover, automatic differentiation alone cannot share repeated calculations of derivatives when we arrive at the same state more than once. 79 | 80 | 81 | Review #87C 82 | =========================================================================== 83 | 84 | Overall merit 85 | ------------- 86 | 87 | B. I support accepting this paper but will not champion it. 88 | 89 | Reviewer expertise 90 | ------------------ 91 | 92 | Y. Knowledgeable 93 | 94 | Paper summary 95 | ------------- 96 | 97 | The main content of the paper is a formalisation of languages and derivatives. The starting point of the formalisation is the choice of type theory (Agda) instead of set theory. Languages are taken to be type-level predicates (A* -> Set). This allows for language membership to have more than just a single justification and to use constructions on types as the logical language to specify language membership. The specification of language membership (language operations) is non-operational and it seems to be one of the key ideas that the authors are advocating, i.e., to have as simple as possible specifications with respect to which other operations (or implementations) can be proved correct. The result of this approach is that the provided language membership deciders (defined using differentiation on inductive or coinductive representations of the language) are correct. In other words, the function that decides language membership is a theorem saying that if there is a representation for a language P, then language membership in P is decidable, i.e., for any w, the function either computes a witness of P w (the type of language membership proofs of w in P) or constructs a proof that no such exists. The rest of the paper works out a generalisation of languages to predicates and predicate operations, investigates some properties of these and how to transport properties to implementations. 98 | 99 | The Agda code did not parse with my version of Agda. Since there was no information of which version of Agda I should use, I did not verify/type check the Agda code. 100 | 101 | Overall, I think this is a nice, clean, well-written and compact paper. Working things out so that clean and simple definitions are also "magically" correct according to Agda is often a non-trivial task. The idea that specifications should be as simple as possible is a very reasonable one. Therefore, I think these ideas deserve to be communicated. 102 | 103 | Comments for author 104 | ------------------- 105 | 106 | * One issue I have is that reading the paper might give the impression that the choice of type theory (instead of set theory) or type-level languages is necessary for achieving non-operational specification of language membership or more informative language membership than just true/false. The paper does demonstrate that type theory can be convenient for achieving this. 107 | 108 | For example, the concatenation of P and Q (subsets of A*) is a certain subset of A* 109 | 110 | P * Q = { w \in A* | exists u, v. w = uv \and u \in P \and v \in Q } 111 | 112 | and the characteristic function of this set corresponds to what is given in Figure 1. This could be done in Agda with decidable languages (A* -> 2) and thus also achieve non-operational specification (values of A* -> 2 as type indices). (Assumption: a word w can be decomposed in a finite number of ways.) 113 | 114 | * Instead of considering languages as A* -> 2 we could take A* -> R where R is a semiring (for example, integer semiring to consider languages as multisets of words instead of just sets of words, i.e., a word occurs in a language some number of times). 115 | 116 | * My main concern regarding this membership question is that this paper is very much related to weighted automata (a subfield of automata theory and formal languages) but there is no discussion of that. The main idea there is exactly that language membership is not necessarily given by (coefficients in) a Boolean semiring. Brzozowski's derivatives have been considered in that setting (Lombardy, Sakarovitch, Derivatives of rational expressions with multiplicity, 2005) and there are many similarities. 117 | 118 | * line 18: I think it would be helpful to later mention what properties would have required induction/coinduction? 119 | 120 | * Contribution 1, line 93: Is this meant to say that capturing the semantic essence of Brzozowski's derivatives is something new? The semantic operation was defined before the syntactic one. Brzozowski references three papers where (under different names) derivatives of sequences are used and then defines the corresponding operation on syntax. 121 | 122 | * Contribution 2, line 99: The duality is between regular expressions and a certain subset of tries (those constructed using the language building operations). The two things named Lang in Figures 4 and 6 are not equally expressive. 123 | 124 | * line 103: Is "proof isomorphism" (throughout the paper) the same as "type isomorphism" on line 188? 125 | 126 | * Contribution 3, line 103: Is this statement considering previous work on formalisations or formal languages in general? As soon as something more general is considered instead of nullability (true/false empty word property), the lemmas (or definitions) for concatenation and closure start to look like those in Figure 2. 127 | 128 | * line 132: Another way to interpret these equations is that the derivative operation is a right monoid action (this might be equivalent to Footnote 4). Furthermore, since we are dealing with free monoids, the right monoid action is given by the generators (i.e., as a left fold). 129 | 130 | * line 179: What is meant exactly by "liberation"? I think it is fairly standard to consider left and right quotients (wrt. a word) of languages (sets of strings, no grammars) as monoid actions and thus reducing language membership to nullability also in a step-by-step manner. 131 | 132 | * line 283: I think it would be good to also say what is the correctness condition (in addition to the fact that it is guaranteed by types). 133 | 134 | * line 308: This left triangle looks like a map operation. Maybe a comment about why this has to be part of the syntax. 135 | 136 | * line 333: Maybe a comment about why f works on the left hand side of the triangle with both p and with p differentiated. 137 | 138 | * Figures 4 and 5: there is no definition of the language (A* -> Set) denoted by an expression and thus no formal statement that it indeed denotes the language in its type. 139 | 140 | * line 462: It is not clear to me why (just) the trie representation exploits the sharing potential. Due to the use of fold, it does not seem to be necessary to use the letter-by-letter differentiation. If f : A* -> B and u : A* then as automatic differentiation take \nu f', f' where f' = D f u. The same pattern can be used with symbolic differentiation by extending it to words via a fold. This part of the paper was a bit underwhelming for me. I was curious to learn what is automatic differentiation. It was described as some reinterpretation of function building vocabulary that avoids duplicated work. From this example it seems that everything here can avoid duplicate work. 141 | 142 | * line 527: "booleans" is sometimes required to be capitalised. 143 | 144 | * line 580: Maybe an example of a property would be nice here. 145 | 146 | * line 623: By a choice of an equivalence relation we are quotienting the (encapsulated) representation. It would be good to give a bit more intuition here (which quotient is this). 147 | 148 | * line 722: Antimirov's derivatives, for example, have also been considered as an inductively defined relation. Then we can take the derivation trees for (e, w, e') as the parsings or explanations (where e is an expression e, w is a word, e' is a nullable expression). 149 | 150 | * line 723: What are transformations? 151 | 152 | * line 724: "based *on* the"? 153 | 154 | 155 | Review #87D 156 | =========================================================================== 157 | 158 | Overall merit 159 | ------------- 160 | 161 | A. I will champion accepting this paper. 162 | 163 | Reviewer expertise 164 | ------------------ 165 | 166 | X. Expert 167 | 168 | Paper summary 169 | ------------- 170 | 171 | The authors present a simple and straightforward formalisation of the basics of formal languages using Brzozowski derivatives in Agda. 172 | 173 | Comments for author 174 | ------------------- 175 | 176 | 177 | # Overview 178 | 179 | Overall this is an eminently readable paper presenting interesting work and, possibly, an interesting addition to the Agda standard library, and I would happily recommend its acceptance based on that alone. 180 | 181 | ## Major comments 182 | 183 | Section 6.2 is quite dense, and it's motivation is not entirely clear to me. Is it's purpose to show that existentials can be used to force the dependently-typed definitions developed throughout the paper into the simply-typed algebraic structures? If so, I believe a single paragraph would to, and would perhaps be less confusing. Furthermore, the section title is confusing: transport properties to what? We haven't defined the existentially quantifier variants yet, so it's not yet obvious that we have anywhere to transport properties to. 184 | 185 | 186 | ## Minor comments 187 | 188 | The title of this paper could be made more explicit about its content: more common use of the term "differentiation" refers to differentiation of real-valued functions (as known in Calculus or Machine Learning), and so the title may sound ambiguous. 189 | 190 | * Line 86: The clause "including the use of existential quantifiers" is confusing, as it in no way makes clear what the existential quantifiers are used for. I'd suggest dropping the clause entirely. 191 | 192 | * Line 125: Please spend a few more words giving an intuition for the derivative 𝒟. 193 | 194 | * Figs. 5 and 7: The column-major versus row-major distinction is rather vague, and requires a bit more explanation of the distinction between data and codata. For instance, w.r.t. the statement "each row is a complete definition" (line 416), it is difficult for the reader unfamiliar with copatterns to infer that, e.g., the first row contains a definition of the actions of \nu and \delta on \nothing, rather than an incomplete definition of \nu and \delta (as it would be with patterns). 195 | 196 | * Line 500: The types of the primitive language constructs are elided throughout the paper. Therefore, it is a bit surprising to see that the variant of star which, ostensibly, abstracts away from lists, still references the list primitives foldr and All, as well as the monoid properties. 197 | 198 | * Line 526: The closure operation is typeset with a different symbol. Is this intentional? Is it meant to refer the 199 | 200 | * Line 568: The authors include an interpretation of the algebraic structure instantiated to natural numbers (as number of parses). However, they do not give a similar interpretation for probability distributions, and they do not motivate how tropical semirings would help with optimisation. Such additions would be very welcome! 201 | 202 | * Line 608: The references to Figures 1 and 4 are rather opaque, especially given that these name are overloaded throughout the paper. Please mention in the text whether these refer to the inductive or the coinductive variant, to save the reader the long journey back to these figures. 203 | 204 | * Line 608: Why are the existentially-wrapped versions specifically based on the inductive variant? Is there anything specific that makes these better or more amenable? The text so far led me to believe the coinductive variant was better and more performant. 205 | 206 | * Line 619: The use of inj2 is confusing to readers familiar with the Agda standard library, as it usually refers to the right injection for sum types. 207 | 208 | * Line 638: Is the pattern matching in the type valid Agda syntax? Again, I don't think these liftings really add anything to the paper. 209 | 210 | * Fig. 14: I don't think the examples of opening various algebraic structures from the standard library adds anything to this paper, and if anything makes it more confusing. Neither the referred algebraic structures, nor the modules Symbolic and Automatic were define in this paper. I think removing this figure is probably preferable over adding explanations of the various structures. 211 | 212 | * * * * * * * * * * * * * * * * * * * * 213 | 214 | Comment @A1 by Reviewer A 215 | --------------------------------------------------------------------------- 216 | The reviewers want to thank the author for this interesting submission and are happy to accept it with the understanding that the author will make all of the following changes: 217 | 218 | 1. Clarify what's new in the context of proof mechanization or type theory, what's new in the context of parsing, and what's new in the context of formal languages & automata. In particular, some contributions do not seem to be novel in the last context. The author did not respond to Reviewer C's comment about the work of Lombardy and Sakarovitch, but their work seems to be similar to the derivative operations in the paper and the author should pay attention to it (in addition to other related works). 219 | 2. Explain the Agda syntax more for the readers who are not familiar with Agda. For example, add concrete examples showing how operators in Fig. 1 apply to concrete regular expressions and words, including the fully expanded types and values (as witnesses of the membership). The current draft has only 16 pages so space should not be an issue. 220 | 3. Clarify the difference between data and codata: both generally and in terms of difference it makes for main results of this paper. 221 | 4. Consider explaining how to recover the ordinary setting where a language is a set of words and language union is idempotent. For example, explain how the current work can be used to prove the classical theorem for the ordinary setting that the set of derivatives of a regular expression is finite modulo associativity, commutativity and idempotence of union (although the derivative defined here might require more). 222 | 5. Fix typos and other issues that have already been discussed in the reviews and the response. 223 | -------------------------------------------------------------------------------- /Predicate/Properties.lagda: -------------------------------------------------------------------------------- 1 | \section{Predicate Properties} 2 | 3 | \begin{code} 4 | {-# OPTIONS --safe #-} 5 | 6 | -- Some properties of predicates 7 | 8 | open import Level 9 | 10 | module Predicate.Properties {ℓ : Level} where 11 | 12 | open import Data.Unit.Polymorphic using (⊤) 13 | open import Data.Product 14 | open import Data.Product.Properties using (×-≡,≡↔≡; ∃∃↔∃∃) 15 | open import Data.Product.Algebra 16 | open import Data.Sum using (inj₁; inj₂) 17 | open import Algebra.Core 18 | open import Algebra.Definitions 19 | open import Algebra.Bundles 20 | open import Algebra.Module.Bundles 21 | open import Relation.Binary.PropositionalEquality 22 | using (_≡_; _≗_; cong; module ≡-Reasoning) renaming (refl to refl≡; sym to sym≡) 23 | open import Function using (id; _∘_; const; flip; _↔_; Inverse; mk↔′) 24 | open import Relation.Unary using (_⊢_) 25 | 26 | open import Data.Product.Algebra using (Σ-assoc) 27 | 28 | -- Local 29 | open import Misc {ℓ} 30 | open import Inverses {ℓ} -- using (module ↔Eq; module ↔R; ∃≡ˡ) 31 | import Algebra.Construct.Function as C 32 | open import Closed ; open Closed.Types {ℓ} 33 | 34 | open ClosedCommutativeSemiring ×-⊎-closedCommutativeSemiring 35 | 36 | -- Pred as semimodule 37 | Pred-Semimodule : Set ℓ → Semimodule (×-⊎-commutativeSemiring ℓ) _ _ 38 | Pred-Semimodule A = C.semimodule A 39 | 40 | module Pred-Semimodule {A} = Semimodule (Pred-Semimodule A) 41 | open Pred-Semimodule 42 | 43 | open import Predicate {ℓ} 44 | 45 | -- Most of these attempts fail in the presence of level polymorphism, because 46 | -- equational reasoning and sym require before & after to have the same type, 47 | -- while type isomorphisms often change levels. 48 | 49 | private 50 | variable 51 | A B C D : Set ℓ 52 | a b c d : A 53 | p q r s : Set ℓ 54 | P Q R S : Pred A 55 | 56 | 57 | infixr 2 _↦_ 58 | _↦_ : A → Set ℓ → Pred A 59 | (a ↦ b) w = (w ≡ a) × b 60 | 61 | infixr 2 _⤇_ -- or ⇰ 62 | _⤇_ : (B → A) → (B → Set ℓ) → B → Pred A 63 | (f ⤇ g) a = f a ↦ g a 64 | 65 | infixr 2 _⇐_ 66 | _⇐_ : (A → B) → (Pred A → Pred B) 67 | _⇐_ = mapⱽ 68 | 69 | pred-decomp : P ⟷ (∃¹ λ a → a ↦ P a) 70 | -- pred-decomp = sym (trans (∃-cong (×-congʳ ≡↔)) ∃≡ˡ) where open ↔Eq 71 | 72 | pred-decomp {P = P} {w} = 73 | begin 74 | P w 75 | ≈˘⟨ ∃≡ˡ ⟩ 76 | (∃ λ a → a ≡ w × P a) 77 | ≈⟨ ∃-cong (×-congʳ sym↔) ⟩ 78 | (∃ λ a → w ≡ a × P a) 79 | ≡⟨⟩ 80 | (∃ λ a → (a ↦ P a) w) 81 | ≡⟨⟩ 82 | (∃¹ λ a → a ↦ P a) w 83 | ∎ 84 | where open ↔R 85 | 86 | -- TODO: rename ∃¹ to ∑ᴹ 87 | 88 | pureⱽ-cong : a ≡ b → pureⱽ a ⟷ pureⱽ b 89 | pureⱽ-cong refl≡ = ↔Eq.refl 90 | 91 | pureⱽ, : ∀ {a : A}{b : B} → pureⱽ (a , b) ⟷ (pureⱽ a ⟨×⟩ pureⱽ b) 92 | pureⱽ, = ↔Eq.sym ×-≡,≡↔≡ 93 | 94 | -- pureⱽ, {a = a}{b = b}{w = (u , v)} = let open ↔R in 95 | -- begin 96 | -- pureⱽ (a , b) (u , v) 97 | -- ≡⟨⟩ 98 | -- (u , v) ≡ (a , b) 99 | -- ≈˘⟨ ×-≡,≡↔≡ ⟩ 100 | -- (u ≡ a × v ≡ b) 101 | -- ≡⟨⟩ 102 | -- (pureⱽ a ⟨×⟩ pureⱽ b) (u , v) 103 | -- ∎ 104 | 105 | ↦-cong : a ≡ b → p ↔ q → (a ↦ p) ⟷ (b ↦ q) 106 | ↦-cong refl≡ = *ᵣ-congˡ 107 | 108 | ↦-congˡ : p ↔ q → (a ↦ p) ⟷ (a ↦ q) 109 | ↦-congˡ = ↦-cong refl≡ 110 | 111 | ↦-congʳ : a ≡ b → (a ↦ p) ⟷ (b ↦ p) 112 | ↦-congʳ = flip ↦-cong ↔Eq.refl 113 | 114 | ↦-distrib-+ : (a ↦ p + q) ⟷ ((a ↦ p) +ᴹ (a ↦ q)) 115 | ↦-distrib-+ {a = a}{p}{q} = let open ⟷R in 116 | begin 117 | (a ↦ p + q) 118 | ≡⟨⟩ 119 | pureⱽ a *ᵣ (p + q) 120 | ≈⟨ *ᵣ-distribˡ (pureⱽ a) p q ⟩ 121 | pureⱽ a *ᵣ p +ᴹ pureⱽ a *ᵣ q 122 | ≡⟨⟩ 123 | (a ↦ p) +ᴹ (a ↦ q) 124 | ∎ 125 | 126 | *ᵣ-distrib-∃ˡ : (P *ᵣ ∃ Q) ⟷ ∃¹ (λ u → P *ᵣ Q u) 127 | *ᵣ-distrib-∃ˡ = ×-distrib-∃ˡ 128 | 129 | -- *ᵣ-distrib-∃ˡ {P = P} {Q = Q} {w} = let open ↔R in 130 | -- begin 131 | -- (P *ᵣ ∃ Q) w 132 | -- ≡⟨⟩ 133 | -- P w * ∃ Q 134 | -- ≈⟨ ×-distrib-∃ˡ ⟩ 135 | -- ∃ (λ u → P w * Q u) -- ≡⟨⟩ ∃ (P w ·ₗ Q) 136 | -- ≡⟨⟩ 137 | -- ∃¹ (λ u → P *ᵣ Q u) w 138 | -- ∎ 139 | 140 | -- ↦-distrib-∃ˡ : (a ↦ ∃ Q) ⟷ ∃¹ (λ u → a ↦ Q u) 141 | ↦-distrib-∃ˡ : (a ↦ ∃ Q) ⟷ ∃¹ (const a ⤇ Q) 142 | ↦-distrib-∃ˡ = ×-distrib-∃ˡ 143 | 144 | -- ↦-distrib-∃ˡ {a = a} = *ᵣ-distrib-∃ˡ {P = pureⱽ a} 145 | 146 | -- ↦-distrib-∃ˡ {a = a}{Q = Q} = let open ⟷R in 147 | -- begin 148 | -- (a ↦ ∃ Q) 149 | -- ≡⟨⟩ 150 | -- pureⱽ a *ᵣ ∃ Q 151 | -- ≈⟨ *ᵣ-distrib-∃ˡ {P = pureⱽ a} ⟩ 152 | -- ∃¹ (λ u → pureⱽ a *ᵣ Q u) 153 | -- ≡⟨⟩ 154 | -- ∃¹ (λ u → a ↦ Q u) 155 | -- ∎ 156 | 157 | -- TODO: try replacing explicit extensional equality with implicit. 158 | 159 | mapⱽ-cong : ∀ {f g : A → B} → f ≗ g → P ⟷ Q → mapⱽ f P ⟷ mapⱽ g Q 160 | -- mapⱽ-cong f≗g P⟷Q = ∃-cong (*-cong (trans↔ (f≗g _)) P⟷Q) 161 | 162 | mapⱽ-cong {P = P}{Q} {f = f}{g} f≗g P⟷Q {w} = let open ⟷R in 163 | begin 164 | mapⱽ f P 165 | ≡⟨⟩ 166 | ∃¹ (λ u → f u ↦ P u) 167 | ≈⟨ ∃-cong (↦-cong (f≗g _) P⟷Q) ⟩ 168 | ∃¹ (λ u → g u ↦ Q u) 169 | ≡⟨⟩ 170 | mapⱽ g Q 171 | ∎ 172 | 173 | mapⱽ-congʳ : ∀ {f g : A → B} → f ≗ g → mapⱽ f P ⟷ mapⱽ g P 174 | mapⱽ-congʳ f≗g {w} = mapⱽ-cong f≗g (⟷Eq.refl {w = w}) 175 | 176 | mapⱽ-congˡ : ∀ {f : A → B} → P ⟷ Q → mapⱽ f P ⟷ mapⱽ f Q 177 | mapⱽ-congˡ P⟷Q {w} = mapⱽ-cong (λ _ → refl≡) P⟷Q 178 | 179 | mapⱽ₂-cong : ∀ {f g : A → B → C} → (∀ a b → f a b ≡ g a b) → P ⟷ R → Q ⟷ S 180 | → mapⱽ₂ f P Q ⟷ mapⱽ₂ g R S 181 | mapⱽ₂-cong {A = A}{B = B} f∼g P∼R Q∼S = 182 | mapⱽ-cong (uncurry f∼g) (λ {(a , b)} → ⟨×⟩-cong (P∼R {w = a}) (Q∼S {w = b}) {w = a , b}) 183 | 184 | mapⱽ₂-congˡ : ∀ {f : A → B → C} → P ⟷ R → Q ⟷ S 185 | → mapⱽ₂ f P Q ⟷ mapⱽ₂ f R S 186 | mapⱽ₂-congˡ = mapⱽ₂-cong (λ _ _ → refl≡) 187 | 188 | mapⱽ₂-congʳ : ∀ {f g : A → B → C} → (∀ a b → f a b ≡ g a b) 189 | → mapⱽ₂ f P Q ⟷ mapⱽ₂ g P Q 190 | mapⱽ₂-congʳ f∼g = mapⱽ-congʳ (uncurry f∼g) 191 | -- mapⱽ₂-congʳ f∼g = mapⱽ₂-cong f∼g ⟷Eq.refl ⟷Eq.refl 192 | 193 | -- mapⱽ₃-congʳ : ∀ {f g : A → B → C → D} → (∀ a b c → f a b c ≡ g a b c) 194 | -- → mapⱽ₃ f P Q R ⟷ mapⱽ₃ g P Q R 195 | -- mapⱽ₃-congʳ f∼g = mapⱽ-congʳ (uncurry₃ʳ f∼g) 196 | 197 | -- mapⱽ is a (covariant) functor 198 | 199 | mapⱽ-id : mapⱽ id P ⟷ P 200 | mapⱽ-id {P = P} = 201 | begin 202 | mapⱽ id P 203 | ≡⟨⟩ 204 | (∃¹ λ a → id a ↦ P a) 205 | ≡⟨⟩ 206 | (∃¹ λ a → a ↦ P a) 207 | ≈˘⟨ pred-decomp ⟩ 208 | P 209 | ∎ 210 | where open ⟷R 211 | 212 | -- mapⱽ-id {w = w} = ⟷Eq.sym (pred-decomp {w = w}) {w = w} 213 | 214 | mapⱽ-∘ : ∀ {f : A → B} {g : B → C} → mapⱽ (g ∘ f) P ⟷ mapⱽ g (mapⱽ f P) 215 | mapⱽ-∘ {A = A} {B = B} {C = C}{P = P}{f}{g} {w = w} = 216 | begin 217 | mapⱽ (g ∘ f) P w 218 | ≡⟨⟩ 219 | (∃ λ a → w ≡ g (f a) × P a) 220 | ≈˘⟨ ∃-cong ∃≡² ⟩ 221 | (∃₂ λ a b → w ≡ g b × b ≡ f a × P a) 222 | ≈⟨ ∃∃↔∃∃ _ ⟩ 223 | (∃₂ λ b a → w ≡ g b × b ≡ f a × P a) 224 | ≈⟨ ∃-cong (∃∃↔∃∃ _) ⟩ 225 | (∃ λ b → w ≡ g b × ∃ λ a → b ≡ f a × P a) 226 | ≡⟨⟩ 227 | (∃ λ b → w ≡ g b × mapⱽ f P b) 228 | ≡⟨⟩ 229 | (∃¹ λ b → g b ↦ mapⱽ f P b) w 230 | ≡⟨⟩ 231 | mapⱽ g (mapⱽ f P) w 232 | ∎ 233 | where open ↔R 234 | 235 | -- Try mapⱽ-∘ with ⟷R instead of ↔R (dropping w): 236 | 237 | -- (a ↦ p) *ᵣ q ⟷ (a ↦ p * q) 238 | -- p ·ₗ (a ↦ q) ⟷ (a ↦ p * q) 239 | 240 | ↦-↦ : ∀ {g : B → A} → ∃¹ (λ b′ → g b′ ↦ (b ↦ p) b′) ⟷ (g b ↦ p) 241 | ↦-↦ {b = b}{p = p}{g = g} {w = w} = let open ↔R in 242 | begin 243 | ∃¹ (λ b′ → g b′ ↦ (b ↦ p) b′) w 244 | ≡⟨⟩ 245 | ∃ (λ b′ → (g b′ ↦ (b ↦ p) b′) w) 246 | ≡⟨⟩ 247 | ∃ (λ b′ → (g b′ ↦ pureⱽ b b′ * p) w) 248 | ≡⟨⟩ 249 | ∃ (λ b′ → pureⱽ (g b′) w * (pureⱽ b b′ * p)) 250 | ≈⟨ ∃≡² ⟩ 251 | pureⱽ (g b) w * p 252 | ≡⟨⟩ 253 | (g b ↦ p) w 254 | ∎ 255 | 256 | mapⱽ-∘₂ : ∀ {f : A → B} {g : B → C} → mapⱽ (g ∘ f) P ⟷ mapⱽ g (mapⱽ f P) 257 | mapⱽ-∘₂ {A = A} {B = B} {C = C}{P = P}{f}{g} = let open ⟷R in 258 | begin 259 | mapⱽ (g ∘ f) P 260 | ≡⟨⟩ 261 | ∃¹ (λ a → g (f a) ↦ P a) 262 | ≈˘⟨ ∃-cong ↦-↦ ⟩ 263 | ∃² (λ a b → g b ↦ (f a ↦ P a) b) 264 | ≈⟨ ∃∃↔∃∃ _ ⟩ 265 | ∃² (λ b a → g b ↦ (f a ↦ P a) b) 266 | ≡⟨⟩ 267 | ∃¹ (λ b → ∃¹ λ a → g b ↦ (f a ↦ P a) b) 268 | ≈˘⟨ ∃-cong ↦-distrib-∃ˡ ⟩ 269 | ∃¹ (λ b → g b ↦ ∃ (λ a → (f a ↦ P a) b)) 270 | ≡⟨⟩ 271 | ∃¹ (λ b → g b ↦ ∃¹ (λ a → f a ↦ P a) b) 272 | ≡⟨⟩ 273 | ∃¹ (λ b → g b ↦ ∃¹ (f ⤇ P) b) 274 | ≡⟨⟩ 275 | ∃¹ (λ b → g b ↦ mapⱽ f P b) 276 | ≡⟨⟩ 277 | ∃¹ (g ⤇ mapⱽ f P) 278 | ≡⟨⟩ 279 | mapⱽ g (mapⱽ f P) 280 | ∎ 281 | 282 | mapⱽ-⁻¹ : ∀ {f : A → B} {g : B → A} → (g ∘ f ≗ id) → mapⱽ g (mapⱽ f P) ⟷ P 283 | mapⱽ-⁻¹ {P = P}{f = f}{g = g} g∘f≗id = let open ⟷R in 284 | begin 285 | mapⱽ g (mapⱽ f P) 286 | ≈˘⟨ mapⱽ-∘ ⟩ 287 | mapⱽ (g ∘ f) P 288 | ≈⟨ mapⱽ-congʳ g∘f≗id ⟩ 289 | mapⱽ id P 290 | ≈⟨ mapⱽ-id ⟩ 291 | P 292 | ∎ 293 | 294 | -- Note similarity to the proof that 𝒟 (f⁻¹) ≡ (𝒟 f)⁻¹ via the chain rule. 295 | 296 | -- TODO: define _⟺_ or similar for function-lifted _⟷_. Use in mapⱽ-id, mapⱽ-∘, 297 | -- and mapⱽ-⁻¹, eg mapⱽ g ∘ mapⱽ f ⟺ id . 298 | 299 | -- TODO: Tidy mapⱽ-id and mapⱽ-∘ proofs using _⤇_ . 300 | 301 | -- TODO: can we prove mapⱽ-∘ without introducing w? I tried and failed. 302 | 303 | -- *ᵣ⟨×⟩*ᵣ : (P *ᵣ p ⟨×⟩ Q *ᵣ q) ⟷ (P ⟨×⟩ Q) *ᵣ (p × q) 304 | -- *ᵣ⟨×⟩*ᵣ = ×-transpose 305 | 306 | -- *ᵣ⟨×⟩*ᵣ {P = P}{p = p}{Q = Q}{q = q} {w = u , v} = let open ↔R in 307 | -- begin 308 | -- (P *ᵣ p ⟨×⟩ Q *ᵣ q) (u , v) 309 | -- ≡⟨⟩ 310 | -- ((P *ᵣ p) u × (Q *ᵣ q) v) 311 | -- ≡⟨⟩ 312 | -- ((P u × p) × (Q v × q)) 313 | -- ≈⟨ ×-transpose ⟩ 314 | -- ((P u × Q v) × (p × q)) 315 | -- ≡⟨⟩ 316 | -- ((P ⟨×⟩ Q) *ᵣ (p × q)) (u , v) 317 | -- ∎ 318 | 319 | ,↦× : {a : A} {b : B} {p q : Set ℓ} → ((a , b) ↦ (p × q)) ⟷ ((a ↦ p) ⟨×⟩ (b ↦ q)) 320 | ,↦× {a = a}{b}{p}{q} = let open ⟷R in 321 | begin 322 | ((a , b) ↦ p × q) 323 | ≡⟨⟩ 324 | pureⱽ (a , b) *ᵣ (p × q) 325 | ≈⟨ *ᵣ-cong {x = pureⱽ _} pureⱽ, ↔Eq.refl ⟩ 326 | (pureⱽ a ⟨×⟩ pureⱽ b) *ᵣ (p × q) 327 | ≈⟨ ×-transpose ⟩ 328 | ((pureⱽ a *ᵣ p) ⟨×⟩ (pureⱽ b *ᵣ q)) 329 | ≡⟨⟩ 330 | ((a ↦ p) ⟨×⟩ (b ↦ q)) 331 | ∎ 332 | 333 | -- TODO: eliminate uses of ×-transpose, which is proved via ×-comm. 334 | -- Maybe pureⱽ generally do commute because their values are 1 or 0. 335 | 336 | -- TODO: I think _⟨×⟩_ is a direct product of semimodules. What about semidirect 337 | -- products? IIRC, Brent Yorgey and Dan Piponi have each done something with 338 | -- semidirect products. 339 | 340 | infixr 2 _⊗_ 341 | _⊗_ : (A → C) → (B → D) → (A × B → C × D) 342 | (f ⊗ g) (a , b) = (f a , g b) 343 | 344 | mapⱽ-⊗ : ∀ {f : A → C} {g : B → D} 345 | → mapⱽ (f ⊗ g) (P ⟨×⟩ Q) ⟷ (mapⱽ f P ⟨×⟩ mapⱽ g Q) 346 | mapⱽ-⊗ {P = P}{Q = Q}{f = f}{g} = let open ⟷R in 347 | begin 348 | ∃¹ ((f ⊗ g) ⤇ (P ⟨×⟩ Q)) 349 | ≈⟨ ∃¹-cong {f = _ ⤇ _} ,↦× ⟩ 350 | ∃¹ ((f ⤇ P) ⟪×⟫ (g ⤇ Q)) 351 | ≈⟨ ∃-distrib-⟨×⟩ ⟩ 352 | (∃¹ (f ⤇ P) ⟨×⟩ ∃¹ (g ⤇ Q)) 353 | ∎ 354 | 355 | mapⱽ-map₁ : ∀ {f : A → C} → mapⱽ (map₁ f) (P ⟨×⟩ Q) ⟷ (mapⱽ f P ⟨×⟩ Q) 356 | mapⱽ-map₁ {P = P}{Q = Q}{f = f} = let open ⟷R in 357 | begin 358 | mapⱽ (map₁ f) (P ⟨×⟩ Q) 359 | ≡⟨⟩ 360 | mapⱽ (f ⊗ id) (P ⟨×⟩ Q) 361 | ≈⟨ mapⱽ-⊗ ⟩ 362 | (mapⱽ f P ⟨×⟩ mapⱽ id Q) 363 | ≈⟨ *-congˡ mapⱽ-id ⟩ 364 | (mapⱽ f P ⟨×⟩ Q) 365 | ∎ 366 | 367 | mapⱽ-map₂ : ∀ {g : B → D} → mapⱽ (map₂ g) (P ⟨×⟩ Q) ⟷ (P ⟨×⟩ mapⱽ g Q) 368 | mapⱽ-map₂ {P = P}{Q = Q}{g = g} = let open ⟷R in 369 | begin 370 | mapⱽ (map₂ g) (P ⟨×⟩ Q) 371 | ≡⟨⟩ 372 | mapⱽ (id ⊗ g) (P ⟨×⟩ Q) 373 | ≈⟨ mapⱽ-⊗ ⟩ 374 | (mapⱽ id P ⟨×⟩ mapⱽ g Q) 375 | ≈⟨ *-congʳ mapⱽ-id ⟩ 376 | (P ⟨×⟩ mapⱽ g Q) 377 | ∎ 378 | 379 | -- Connection between covariant and contravariant functors 380 | 381 | module _ (A↔B : A ↔ B) where 382 | open Inverse A↔B 383 | -- TODO: Can we avoid K? 384 | open import Axiom.UniquenessOfIdentityProofs.WithK using (uip) 385 | 386 | ≡⁻¹ : b ≡ f a ↔ a ≡ f⁻¹ b 387 | ≡⁻¹ {b = b}{a = a} = mk↔′ 388 | (λ {refl≡ → sym≡ (inverseʳ a)}) 389 | (λ {refl≡ → sym≡ (inverseˡ b)}) 390 | (λ a≡f⁻¹b → uip _ a≡f⁻¹b) 391 | (λ a≡f⁻¹b → uip _ a≡f⁻¹b) 392 | 393 | -- ⇐-⊢ : ∀ {A↔B : A ↔ B} → mapⱽ f (P ∘ f) ⟷ P 394 | ⇐-⊢ : (f ⇐ (f ⊢ P)) ⟷ P 395 | ⇐-⊢ {P = P} = let open ⟷R in 396 | begin 397 | mapⱽ f (P ∘ f) 398 | ≡⟨⟩ 399 | ∃¹ (f ⤇ P ∘ f) 400 | ≡⟨⟩ 401 | ∃¹ (λ a → f a ↦ P (f a)) 402 | ≡⟨⟩ 403 | ∃¹ (λ a → pureⱽ (f a) *ᵣ P (f a)) 404 | ≡⟨⟩ 405 | ∃¹ (λ a b → b ≡ f a × P (f a)) 406 | ≈⟨ ∃¹-cong (λ {a b} → *-congʳ (≡⁻¹ {b = b})) ⟩ 407 | ∃¹ (λ a b → a ≡ f⁻¹ b × P (f a)) 408 | ≡⟨⟩ 409 | (λ b → ∃ λ a → a ≡ f⁻¹ b × P (f a)) 410 | ≈⟨ ∃≡ˡ ⟩ 411 | (λ b → P (f (f⁻¹ b))) 412 | ≈⟨ ≡↔ (cong P (inverseˡ _)) ⟩ 413 | (λ b → P b) 414 | ≡⟨⟩ 415 | P 416 | ∎ 417 | 418 | -- (f ⇐ (f ⊢ P)) ⟷ P 419 | -- f ⇐_ ∘ f ⊢_ ⟺ id 420 | 421 | -- TODO: Something about g ∘ f ≡ id → (g ⇐ P) ⟷ (f ⊢ P) 422 | -- Rethink the operator names. 423 | 424 | mapⱽ-comm : ∀ {A : Set ℓ} {_∙_ : Op₂ A} → Commutative _≡_ _∙_ → Commutative _⟷_ (mapⱽ₂ _∙_) 425 | mapⱽ-comm {A = A}{_∙_ = _∙_} ∙-comm P Q = let open ⟷R in 426 | begin 427 | mapⱽ ◻ (P ⟨×⟩ Q) 428 | ≈⟨ mapⱽ-cong (uncurry ∙-comm) (*-comm _ _) ⟩ -- ⟨×⟩-comm {P = P} {Q = Q} 429 | mapⱽ (◻ ∘ swap) ((Q ⟨×⟩ P) ∘ swap) 430 | ≈⟨ mapⱽ-∘ ⟩ 431 | mapⱽ ◻ (mapⱽ swap ((Q ⟨×⟩ P) ∘ swap)) 432 | ≈⟨ mapⱽ-congˡ (⇐-⊢ (*-comm A A)) ⟩ 433 | mapⱽ ◻ (Q ⟨×⟩ P) 434 | ∎ 435 | where infixl 7 _⋆_ 436 | _⋆_ = mapⱽ₂ _∙_ 437 | ◻ = uncurry _∙_ 438 | 439 | -- TODO: consider using converse from Relation.Unary: 440 | -- 441 | -- _~ : Pred (A × B) ℓ → Pred (B × A) ℓ 442 | -- P ~ = P ∘ swap 443 | 444 | mapⱽ-assoc : ∀ {A : Set ℓ} {∙ : Op₂ A} → Associative _≡_ ∙ → Associative _⟷_ (mapⱽ₂ ∙) 445 | mapⱽ-assoc {A = A}{∙ = ∙} ∙-assoc P Q R = let open ⟷R in 446 | begin 447 | (P ⋆ Q) ⋆ R 448 | ≡⟨⟩ 449 | mapⱽ₂ ∙ (mapⱽ₂ ∙ P Q) R 450 | ≡⟨⟩ 451 | mapⱽ ◻ (mapⱽ ◻ (P ⟨×⟩ Q) ⟨×⟩ R) 452 | ≈˘⟨ mapⱽ-congˡ mapⱽ-map₁ ⟩ 453 | mapⱽ ◻ (mapⱽ (map₁ ◻) ((P ⟨×⟩ Q) ⟨×⟩ R)) 454 | ≈˘⟨ mapⱽ-∘ ⟩ 455 | mapⱽ (◻ ∘ map₁ ◻) ((P ⟨×⟩ Q) ⟨×⟩ R) 456 | ≈⟨ mapⱽ-cong (uncurry₃ˡ ∙-assoc) (*-assoc _ _ _) ⟩ 457 | mapⱽ (◻ ∘ map₂ ◻ ∘ assocʳ) ((P ⟨×⟩ (Q ⟨×⟩ R)) ∘ assocʳ) 458 | ≈⟨ mapⱽ-∘ ⟩ 459 | mapⱽ (◻ ∘ map₂ ◻) (mapⱽ assocʳ ((P ⟨×⟩ (Q ⟨×⟩ R)) ∘ assocʳ)) 460 | ≈⟨ mapⱽ-congˡ (⇐-⊢ (*-assoc A A A)) ⟩ 461 | mapⱽ (◻ ∘ map₂ ◻) (P ⟨×⟩ (Q ⟨×⟩ R)) 462 | ≈⟨ mapⱽ-∘ ⟩ 463 | mapⱽ ◻ (mapⱽ (map₂ ◻) (P ⟨×⟩ (Q ⟨×⟩ R))) 464 | ≈⟨ mapⱽ-congˡ mapⱽ-map₂ ⟩ 465 | mapⱽ ◻ (P ⟨×⟩ mapⱽ ◻ (Q ⟨×⟩ R)) 466 | ≡⟨⟩ 467 | mapⱽ₂ ∙ P (mapⱽ₂ ∙ Q R) 468 | ≡⟨⟩ 469 | P ⋆ (Q ⋆ R) 470 | ∎ 471 | where infixl 7 _⋆_ 472 | _⋆_ = mapⱽ₂ ∙ 473 | ◻ = uncurry ∙ 474 | 475 | -- TODO: Simplify mapⱽ-comm and mapⱽ-assoc 476 | 477 | mapⱽ-identityˡ : ∀ {A : Set ℓ} {∙ : Op₂ A} {ε : A} 478 | → LeftIdentity _≡_ ε ∙ → LeftIdentity _⟷_ (pureⱽ ε) (mapⱽ₂ ∙) 479 | mapⱽ-identityˡ {∙ = _∙_}{ε = ε} identityˡ Q {w} = let open ↔R in 480 | begin 481 | (pureⱽ ε ⋆ Q) w 482 | ≡⟨⟩ 483 | (∃ λ (u , v) → w ≡ u ∙ v × pureⱽ ε u × Q v) 484 | ≈⟨ Σ-assoc ⟩ 485 | (∃₂ λ u v → w ≡ u ∙ v × u ≡ ε × Q v) 486 | ≈⟨ ∃∃↔∃∃ _ ⟩ 487 | (∃₂ λ v u → w ≡ u ∙ v × u ≡ ε × Q v) 488 | ≈⟨ ∃-cong ∃≡² ⟩ 489 | (∃ λ v → w ≡ ε ∙ v × Q v) 490 | ≈⟨ ∃-cong (*-congʳ (trans↔ (identityˡ _))) ⟩ 491 | (∃ λ v → w ≡ v × Q v) 492 | ≈⟨ ∃-cong (*-congʳ sym↔) ⟩ 493 | (∃ λ v → v ≡ w × Q v) 494 | ≈⟨ ∃≡ˡ ⟩ 495 | Q w 496 | ∎ 497 | where infixl 7 _⋆_ 498 | _⋆_ = mapⱽ₂ _∙_ 499 | 500 | mapⱽ-identityʳ : ∀ {A : Set ℓ} {∙ : Op₂ A} {ε : A} 501 | → RightIdentity _≡_ ε ∙ → RightIdentity _⟷_ (pureⱽ ε) (mapⱽ₂ ∙) 502 | mapⱽ-identityʳ {∙ = _∙_}{ε = ε} identityʳ P {w} = let open ↔R in 503 | begin 504 | (P ⋆ pureⱽ ε) w 505 | ≡⟨⟩ 506 | (∃ λ (u , v) → w ≡ u ∙ v × P u × pureⱽ ε v) 507 | ≡⟨⟩ 508 | (∃ λ (u , v) → w ≡ u ∙ v × P u × v ≡ ε) 509 | ≈˘⟨ ∃-cong (*-assoc _ _ _) ⟩ 510 | (∃ λ (u , v) → (w ≡ u ∙ v × P u) × v ≡ ε) 511 | ≈⟨ Σ-assoc ⟩ 512 | (∃₂ λ u v → (w ≡ u ∙ v × P u) × v ≡ ε) 513 | ≈⟨ ∃-cong ∃≡ʳ ⟩ 514 | (∃ λ u → w ≡ u ∙ ε × P u) 515 | ≈⟨ ∃-cong (*-congʳ (trans↔ (identityʳ _))) ⟩ 516 | (∃ λ u → w ≡ u × P u) 517 | ≈⟨ ∃-cong (*-congʳ sym↔) ⟩ 518 | (∃ λ u → u ≡ w × P u) 519 | ≈⟨ ∃≡ˡ ⟩ 520 | P w 521 | ∎ 522 | where infixl 7 _⋆_ 523 | _⋆_ = mapⱽ₂ _∙_ 524 | 525 | mapⱽ-identity : ∀ {A : Set ℓ} {∙ : Op₂ A} {ε : A} 526 | → Identity _≡_ ε ∙ → Identity _⟷_ (pureⱽ ε) (mapⱽ₂ ∙) 527 | mapⱽ-identity (identityˡ , identityʳ) = mapⱽ-identityˡ identityˡ , mapⱽ-identityʳ identityʳ 528 | 529 | 530 | mapⱽ-distribˡ : ∀ {A : Set ℓ} {_∙_ : Op₂ A} → _DistributesOverˡ_ _⟷_ (mapⱽ₂ _∙_) _+ᴹ_ 531 | mapⱽ-distribˡ {_∙_ = _∙_} P Q R = let open ⟷R in 532 | begin 533 | P ⋆ (Q +ᴹ R) 534 | ≡⟨⟩ 535 | ∃¹ (λ (u , v) → u ∙ v ↦ P u * (Q v + R v)) 536 | ≈⟨ ∃-cong (↦-congˡ (distribˡ _ _ _)) ⟩ 537 | ∃¹ (λ (u , v) → u ∙ v ↦ (P u * Q v + P u * R v)) 538 | ≈⟨ ∃-cong ↦-distrib-+ ⟩ 539 | ∃¹ (λ (u , v) → (u ∙ v ↦ P u * Q v) +ᴹ (u ∙ v ↦ P u * R v)) 540 | ≈⟨ ∃-distrib-∪ ⟩ 541 | P ⋆ Q +ᴹ P ⋆ R 542 | ∎ 543 | where infixl 7 _⋆_ 544 | _⋆_ = mapⱽ₂ _∙_ 545 | 546 | mapⱽ-distribʳ : ∀ {A : Set ℓ} {_∙_ : Op₂ A} → _DistributesOverʳ_ _⟷_ (mapⱽ₂ _∙_) _+ᴹ_ 547 | mapⱽ-distribʳ {_∙_ = _∙_} R P Q = let open ⟷R in 548 | begin 549 | (P +ᴹ Q) ⋆ R 550 | ≡⟨⟩ 551 | ∃¹ (λ (u , v) → u ∙ v ↦ (P u + Q u) * R v) 552 | ≈⟨ ∃-cong (↦-congˡ (distribʳ _ _ _)) ⟩ 553 | ∃¹ (λ (u , v) → u ∙ v ↦ (P u * R v + Q u * R v)) 554 | ≈⟨ ∃-cong ↦-distrib-+ ⟩ 555 | ∃¹ (λ (u , v) → (u ∙ v ↦ P u * R v) +ᴹ (u ∙ v ↦ Q u * R v)) 556 | ≈⟨ ∃-distrib-∪ ⟩ 557 | P ⋆ R +ᴹ Q ⋆ R 558 | ∎ 559 | where infixl 7 _⋆_ 560 | _⋆_ = mapⱽ₂ _∙_ 561 | 562 | mapⱽ-distrib : ∀ {A : Set ℓ} {_∙_ : Op₂ A} → _DistributesOver_ _⟷_ (mapⱽ₂ _∙_) _+ᴹ_ 563 | mapⱽ-distrib = mapⱽ-distribˡ , mapⱽ-distribʳ 564 | 565 | ↦-0# : (a ↦ 0#) ⟷ 0ᴹ 566 | ↦-0# = *ᵣ-zeroʳ (pureⱽ _) 567 | 568 | -- ↦-0# {a = a} = let open ⟷R in 569 | -- begin 570 | -- (a ↦ 0#) 571 | -- ≡⟨⟩ 572 | -- pureⱽ a *ᵣ 0# 573 | -- ≈⟨ *ᵣ-zeroʳ (pureⱽ a) ⟩ 574 | -- 0ᴹ 575 | -- ∎ 576 | 577 | mapⱽ-zeroˡ : ∀ {A : Set ℓ} {_∙_ : Op₂ A} → LeftZero _⟷_ 0ᴹ (mapⱽ₂ _∙_) 578 | mapⱽ-zeroˡ {A = A}{_∙_ = _∙_} P = let open ⟷R in 579 | begin 580 | 0ᴹ ⋆ P 581 | ≡⟨⟩ 582 | ∃¹ (λ (u , v) → u ∙ v ↦ 0ᴹ u * P v) 583 | ≈⟨ ∃-cong (↦-congˡ (zeroˡ _)) ⟩ 584 | ∃¹ (λ (u , v) → u ∙ v ↦ 0#) 585 | ≈⟨ ∃-cong ↦-0# ⟩ 586 | ∃¹ (const 0ᴹ) 587 | ≡⟨⟩ 588 | const (∃ 0ᴹ) 589 | ≡⟨⟩ 590 | const ((A × A) * 0#) 591 | ≈⟨ zeroʳ (A × A) ⟩ 592 | const 0# 593 | ≡⟨⟩ 594 | 0ᴹ 595 | ∎ 596 | where infixl 7 _⋆_ 597 | _⋆_ = mapⱽ₂ _∙_ 598 | 599 | mapⱽ-zeroʳ : ∀ {A : Set ℓ} {_∙_ : Op₂ A} → RightZero _⟷_ 0ᴹ (mapⱽ₂ _∙_) 600 | mapⱽ-zeroʳ {A = A}{_∙_ = _∙_} P = let open ⟷R in 601 | begin 602 | P ⋆ 0ᴹ 603 | ≡⟨⟩ 604 | ∃¹ (λ (u , v) → u ∙ v ↦ P u * 0ᴹ v) 605 | ≡⟨⟩ 606 | ∃¹ (λ (u , v) → u ∙ v ↦ P u * 0ᴹ v) 607 | ≡⟨⟩ 608 | ∃¹ (λ (u , v) → u ∙ v ↦ P u * 0#) 609 | ≈⟨ ∃-cong (↦-congˡ (zeroʳ _)) ⟩ 610 | ∃¹ (λ (u , v) → u ∙ v ↦ 0#) 611 | ≈⟨ ∃-cong ↦-0# ⟩ 612 | const ((A × A) * 0#) 613 | ≈⟨ zeroʳ (A × A) ⟩ 614 | 0ᴹ 615 | ∎ 616 | where infixl 7 _⋆_ 617 | _⋆_ = mapⱽ₂ _∙_ 618 | 619 | mapⱽ-zero : ∀ {A : Set ℓ} {_∙_ : Op₂ A} → Zero _⟷_ 0ᴹ (mapⱽ₂ _∙_) 620 | mapⱽ-zero = mapⱽ-zeroˡ , mapⱽ-zeroʳ 621 | 622 | import Algebra.Structures as S 623 | 624 | module MonoidSemiringProperties {M : Set ℓ} {_∙_ : Op₂ M} {ε : M} 625 | (isMonoid-M : S.IsMonoid _≡_ _∙_ ε) where 626 | open S.IsMonoid isMonoid-M 627 | 628 | open import Algebra.Structures {suc ℓ} {ℓ} {Pred M} _⟷_ 629 | 630 | open MonoidOps _∙_ ε 631 | 632 | ⋆-cong : P ⟷ R → Q ⟷ S → P ⋆ Q ⟷ R ⋆ S 633 | ⋆-cong = mapⱽ₂-congˡ 634 | 635 | ⋆-congˡ : Q ⟷ S → P ⋆ Q ⟷ P ⋆ S 636 | ⋆-congˡ {P = P} = ⋆-cong (⟷Eq.refl {x = P}) 637 | 638 | ⋆-congʳ : P ⟷ R → P ⋆ Q ⟷ R ⋆ Q 639 | ⋆-congʳ {Q = Q} = flip ⋆-cong (⟷Eq.refl {x = Q}) 640 | 641 | -- ⋆-cong {P}{R}{Q}{S} P⟷R Q⟷S {w} = 642 | -- begin 643 | -- (∃ λ (p , q) → w ≡ p ∙ q × P p × Q q) 644 | -- ≈⟨ ∃-cong (↦-congˡ (*-cong P⟷R Q⟷S)) ⟩ 645 | -- (∃ λ (r , s) → w ≡ r ∙ s × R r × S s) 646 | -- ∎ 647 | -- where open ↔R 648 | -- -- TODO: try again without w 649 | 650 | open import Data.List 651 | open import Data.List.Relation.Unary.All 652 | 653 | ☆-star : P ☆ ⟷ 𝟏 ∪ P ⋆ P ☆ 654 | ☆-star {w = w} = mk↔′ 655 | (λ { ([] , refl≡ , []) → inj₁ refl≡ 656 | ; (p ∷ ps , refl≡ , Pp ∷ Pps) → inj₂ ((p , foldr _∙_ ε ps) , refl≡ , Pp , ps , refl≡ , Pps) }) 657 | (λ { (inj₁ refl≡) → [] , refl≡ , [] 658 | ; (inj₂ ((p , .(foldr _∙_ ε ps)) , refl≡ , Pp , ps , refl≡ , Pps)) → p ∷ ps , refl≡ , Pp ∷ Pps}) 659 | (λ { (inj₁ refl≡) → refl≡ 660 | ; (inj₂ ((p , .(foldr _∙_ ε ps)) , refl≡ , Pp , ps , refl≡ , Pps)) → refl≡ }) 661 | (λ { ([] , refl≡ , []) → refl≡ 662 | ; (p ∷ ps , refl≡ , Pp ∷ Pps) → refl≡ }) 663 | 664 | ✪↔☆ : P ✪ ⟷ P ☆ 665 | ✪↔☆ {P = P} = mk↔′ f f⁻¹ invˡ invʳ 666 | where 667 | f : ∀ {w} → (P ✪) w → (P ☆) w 668 | f zero✪ = [] , refl≡ , [] 669 | f (suc✪ ((u , v) , refl≡ , Pu , P✪v)) with f P✪v 670 | ... | us , refl≡ , Pus = u ∷ us , refl≡ , Pu ∷ Pus 671 | 672 | f⁻¹ : ∀ {w} → (P ☆) w → (P ✪) w 673 | f⁻¹ ([] , refl≡ , []) = zero✪ 674 | f⁻¹ (u ∷ us , refl≡ , Pu ∷ Pus) = 675 | suc✪ ((u , foldr _∙_ ε us) , refl≡ , Pu , f⁻¹ (us , refl≡ , Pus)) 676 | 677 | invˡ : ∀ {w} (z : (P ☆) w) → f (f⁻¹ z) ≡ z 678 | invˡ ([] , refl≡ , []) = refl≡ 679 | invˡ (u ∷ us , refl≡ , Pu ∷ Pus) rewrite invˡ (us , refl≡ , Pus) = refl≡ 680 | 681 | invʳ : ∀ {w} (z : (P ✪) w) → f⁻¹ (f z) ≡ z 682 | invʳ zero✪ = refl≡ 683 | invʳ (suc✪ ((u , v) , refl≡ , Pu , P✪v)) with f P✪v | invʳ P✪v 684 | ... | us , refl≡ , Pus | refl≡ = refl≡ 685 | 686 | {- 687 | 688 | -- Miscellaneous other properties 689 | 690 | ↦-homo-ε : (ε ↦ ⊤) ⟷ 𝟏 691 | ↦-homo-ε = *-identityʳ _ 692 | 693 | -- ↦-homo-ε {w} = 694 | -- begin 695 | -- (ε ↦ ⊤) w 696 | -- ≡⟨⟩ 697 | -- (w ≡ ε × ⊤) 698 | -- ≈⟨ *-identityʳ _ ⟩ 699 | -- w ≡ ε 700 | -- ≡⟨⟩ 701 | -- 𝟏 w 702 | -- ∎ 703 | -- where open ↔R 704 | -- -- open Set-CommutativeSemiring 705 | 706 | ↦-homo-⋆ : ∀ {a b p q} → (a ∙ b ↦ p * q) ⟷ (a ↦ p) ⋆ (b ↦ q) 707 | ↦-homo-⋆ {a}{b}{p}{q} {w} = 708 | begin 709 | (a ∙ b ↦ p * q) w 710 | ≡⟨⟩ 711 | (w ≡ a ∙ b × (p * q)) 712 | ≈˘⟨ ∃≡² ⟩ 713 | (∃ λ (u , v) → w ≡ u ∙ v × (u , v) ≡ (a , b) × (p × q)) 714 | ≈˘⟨ ∃-cong (*-congˡ (*-congʳ ×-≡,≡↔≡)) ⟩ 715 | (∃ λ (u , v) → w ≡ u ∙ v × (u ≡ a × v ≡ b) × (p × q)) 716 | ≈⟨ ∃-cong (*-congˡ ×-transpose) ⟩ 717 | (∃ λ (u , v) → w ≡ u ∙ v × (u ≡ a × p) × (v ≡ b × q)) 718 | ≡⟨⟩ 719 | (∃ λ (u , v) → w ≡ u ∙ v × (a ↦ p) u × (b ↦ q) v) 720 | ≡⟨⟩ 721 | ((a ↦ p) ⋆ (b ↦ q)) w 722 | ∎ 723 | where open ↔R 724 | 725 | -- TODO: Prove homomorphicity of ↦ compositionally from homomorphicity of 726 | -- pureⱽ and *ᵣ. If I can't, then at least prove pureⱽ-homo-⋆ via ↦-homo-⋆. 727 | 728 | ↦1# : (a ↦ 1#) ⟷ pureⱽ a 729 | ↦1# = *ᵣ-identityʳ (pureⱽ _) 730 | 731 | -- ↦1# {a = a} = let open ⟷R in 732 | -- begin 733 | -- (a ↦ 1#) 734 | -- ≡⟨⟩ 735 | -- pureⱽ a *ᵣ 1# 736 | -- ≈⟨ *ᵣ-identityʳ (pureⱽ a) ⟩ 737 | -- pureⱽ a 738 | -- ∎ 739 | 740 | -- pureⱽ-homo-ε holds definitionally. 741 | pureⱽ-homo-ε : pureⱽ ε ⟷ 𝟏 742 | pureⱽ-homo-ε = ↔Eq.refl -- ⟷Eq.refl {x = 𝟏} 743 | 744 | pureⱽ-homo-⋆ : pureⱽ (a ∙ b) ⟷ pureⱽ a ⋆ pureⱽ b 745 | pureⱽ-homo-⋆ {a}{b} = let open ⟷R in 746 | begin 747 | pureⱽ (a ∙ b) 748 | ≈˘⟨ ↦1# ⟩ 749 | (a ∙ b ↦ 1#) 750 | ≈˘⟨ ↦-congˡ (*-identityʳ 1#) ⟩ 751 | (a ∙ b ↦ 1# * 1#) 752 | ≈⟨ ↦-homo-⋆ ⟩ 753 | (a ↦ 1#) ⋆ (b ↦ 1#) 754 | ≈⟨ ⋆-cong ↦1# ↦1# ⟩ 755 | pureⱽ a ⋆ pureⱽ b 756 | ∎ 757 | 758 | -- pureⱽ-homo-⋆ : pureⱽ (a ∙ b) ⟷ pureⱽ a ⋆ pureⱽ b 759 | -- pureⱽ-homo-⋆ {a}{b} {z} = 760 | -- begin 761 | -- z ≡ a ∙ b 762 | -- ≈˘⟨ ∃≡ʳ ⟩ 763 | -- (∃ λ (u , v) → z ≡ u ∙ v × (u , v) ≡ (a , b) ) 764 | -- ≈˘⟨ ∃-cong (*-congˡ ×-≡,≡↔≡) ⟩ 765 | -- (∃ λ (u , v) → z ≡ u ∙ v × u ≡ a × v ≡ b ) 766 | -- ∎ 767 | -- where open ↔R 768 | -} 769 | 770 | \end{code} 771 | -------------------------------------------------------------------------------- /response-to-reviews.md: -------------------------------------------------------------------------------- 1 | # Paper #87 review responses 2 | 3 | ***Note:** I say "we" instead of "I" in these responses, because the blind review process was still in effect when I wrote them.* 4 | 5 | 6 | ## Review A 7 | 8 | > My major concern is novelty, but unfortunately I am not familiar with the prior work. Given that I was a bit surprised by the absence of prior work, and that there were already so many papers on related topics, my rating is B. I might advocate for this paper more enthusiastically once novelty is confirmed. 9 | 10 | We were surprised as well not to find formalized & mechanized connections between the simple, non-operational semantics of language specification and the dual decidable representations of regular expressions and tries. 11 | Note also that proof relevance yields *parsing* (explanations) rather than mere recognition. 12 | We are not aware of formalized (or even rigorous) accounts of *parsing* (as opposed to mere recognition) based on non-operational specification of languages and derivatives. 13 | 14 | > I recognize that a note on the versions of Agda and its standard library would help the reader. 15 | 16 | Will do. Thank you for the suggestion! 17 | 18 | 19 | ## Review B 20 | 21 | > It seems that in Fig. 1 "a language over an alphabet A" is not A^* but A^* -> Set l. This is at least nonstandard --- a standard definition should be A*. Explain the reason. 22 | 23 | `A^*` is the set of *all* strings over an alphabet `A`, i.e., the "universal language". 24 | Other languages are proper subsets, which are neatly formalized in type theory as `A^* -> Set l`, i.e., a function mapping each string `w` to the type of proofs that `w` is in the language. 25 | 26 | > By the way, can you use Prop instead of Set l? I suspect the former is more intuitive, especially for those who are not experts on type theory. 27 | 28 | `Set` appears to be much more widely used in the Agda ecosystem than `Prop` and so using `Set` allows reusing many more existing definitions and proofs, including type isomorphisms. 29 | More fundamentally, `Prop` is proof-irrelevant, so we would only get language recognition. 30 | In contrast, using `Set` gives proof-relevant recognition, i.e., parsing. 31 | 32 | > In Fig. 1, while I believe c :: A, I could not figure out the type of scalar, s. 33 | 34 | In the definition of scalar multiplication (`(s · P) w = s × P w`), `s` appears as an argument to the type product operator and so is a type (i.e., has type `Set ℓ`). 35 | Logically, it's a proposition being conjoined with the property of the string `w` being in the language `P`. 36 | This operator makes for more elegant and more general formulations of derivatives of concatenated and starred languages than Brzozowski's definitions. 37 | The additional generality comes from proof relevance ("parsing"). 38 | It's also fundamentally important to the algebraic setting of semimodules (generalized vector spaces), which is a natural fit to proof-relevant languages. 39 | 40 | > This paper carefully distinguished three kinds of equalities, namely propositional equality, isomorphism, and extensional isomorphism. However, I did not understand why that equality should be used for each occurrence. For example, why in Fig. 1 is propositional equality necessary? 41 | 42 | Figure 1 uses both propositional equality and extensional isomorphism. 43 | We could not have used (propositional) equality for all of the lemmas, since this strong condition does not hold in all cases. 44 | We could have used the weaker isomorphism condition for all of the lemmas, but the equalities are more specific and thus more informative. 45 | Moreover, isomorphisms require some notational overhead to apply, as shown where they are invoked in figures 5 and 7. 46 | In fact, there might be some notational overhead even for propositional equalities, but these ones are actually *definitional* (meaning proved by normalization, as noted on line 194), making them especially convenient. 47 | 48 | The reason for using both isomorphism and extensional isomorphism is simply that the `ν` lemmas are about types/propositions (nullability) while the `δ` lemmas are about predicates (general, proof-relevant language membership). 49 | 50 | *Correction:* My answer was about Figure 3. 51 | As for Figure 1, the uses of propositional equality are for lists (not types or type-level predicates), to which the other mentioned equality notions do not apply. 52 | 53 | > Section 2 uses A* -> B instead of A* -> Set l. Perhaps, the reason is that all the definitions can be polymorphic on the result, B. However, this difference at least makes the meaning difficult to understand. 54 | 55 | Indeed, this generalization is important because it points out that the central notion of derivatives is much more general than its usual setting of languages (even beyond generalizing from `Set` to arbitrary semirings), and some definitions and properties hold in the more general setting. 56 | 57 | > p5. "isomorphic types or predicates are equivalently decidable": 58 | > I could not understand the reason. 59 | 60 | The reason (proof) is outlined in the next sentence ("One direction of the isomorphism proves \AB{B} from \AB{A}, while the other proves {\AF ¬ B} from {\AF ¬ A}") and formalized in the definitions/proofs of the `_▹_` and `_▸_` signatures/lemmas stated in between. 61 | As with the rest of the paper, we mostly state the lemmas & theorems (signatures), leaving the interested reader to view proofs as desired in the paper/project's Agda source code. 62 | 63 | > Figs. 4 and 5 are nearly identical to Figs. 6 and 7. While I could understand that it is because of the duality, I still wonder what is the essential differences. 64 | 65 | The essential difference is that the representation data constructors and the representation examiners trade places (the duality). 66 | More specifically 67 | 68 | * For regular expressions, language building blocks (`∅`, `_∪_`, `𝟙`, `_*_`, etc) are data constructors, while `ν` and `δ` are functions that examine the data. 69 | * For tries, `ν` and `δ` are codata constructors, while language building blocks are functions that examine the codata. 70 | 71 | This duality is visually highlighted in the code arrangement in figures 4 and 6 and in the syntax coloring in figures 5 and 7. 72 | (Line 419: "The compiler-generated syntax coloring differs to reflect the changed interpretation of the definitions.") 73 | 74 | > Usually, an inductive definition gives a datatype, whereas a coinductive one does a function (or stream). However, the automatic differentiation requires the concrete construction process of the function even for the coinductive case. Another possible difference is the potential of the infiniteness of the coinductive definition. However, for this particular case, I am not aware of the usefulness of infiniteness. 75 | 76 | The same pattern occurs in this paper. 77 | Streams are a special case of string tries (for a singleton alphabet), being infinite in the specific and more general case. 78 | The pattern also generalizes beyond string tries to tries over arbitrary regular algebraic data types. 79 | 80 | By "concrete construction process", do you mean that there are still (corecursively defined) examination functions? 81 | If so, this situation is also customary in our experience when working with coinductive structures. 82 | 83 | > On page 9. it is argued that a difference is a work-sharing between the function and its derivative. I agree with this point but do not think it to be critical. Existing parsing methods based on the derivatives do such work sharing. Moreover, automatic differentiation alone cannot share repeated calculations of derivatives when we arrive at the same state more than once. 84 | 85 | We thought this work sharing issue worth pointing out, because it makes an enormous difference for symbolic vs automatic differentiation in the usual calculus sense, and indeed is the main argument usually given in the literature for preferring the latter over the former. 86 | We are not aware of a clear or rigorous discussion of the issue in derivative-based parsing work (though there are many optimizations described informally and without semantic basis). 87 | It is not clear what the practical implications are in the setting of languages, but it seems worth mentioning to stimulate closer examination. 88 | If the performance difference is much less than with differential calculus, then it's important to understand why. 89 | For instance, maybe there is a dual technique for regular expressions that could be examined clearly for useful practical and theoretical insight. 90 | On the other hand, if the difference is the same or greater, then it's important to exploit the potential with better justification and perhaps implementation. 91 | 92 | We certainly agree with your remark about the potential for improved work sharing when a "state" is reached in more than one way! 93 | 94 | 95 | ## Review C 96 | 97 | > The specification of language membership (language operations) is non-operational and it seems to be one of the key ideas that the authors are advocating, i.e., to have as simple as possible specifications with respect to which other operations (or implementations) can be proved correct. 98 | 99 | Yes, exactly. 100 | 101 | > The result of this approach is that the provided language membership deciders (defined using differentiation on inductive or coinductive representations of the language) are correct. In other words, the function that decides language membership is a theorem saying that if there is a representation for a language P, then language membership in P is decidable, i.e., for any w, the function either computes a witness of P w (the type of language membership proofs of w in P) or constructs a proof that no such exists. 102 | 103 | Indeed. Well-summarized. Thank you. 104 | 105 | > Overall, I think this is a nice, clean, well-written and compact paper. Working things out so that clean and simple definitions are also "magically" correct according to Agda is often a non-trivial task. The idea that specifications should be as simple as possible is a very reasonable one. Therefore, I think these ideas deserve to be communicated. 106 | 107 | Thank you! 108 | The clean, simple, automatically verified nature is exactly what we were wanting to highlight. 109 | 110 | > One issue I have is that reading the paper might give the impression that the choice of type theory (instead of set theory) or type-level languages is necessary for achieving non-operational specification of language membership or more informative language membership than just true/false. >The paper does demonstrate that type theory can be convenient for achieving this. 111 | > 112 | > For example, the concatenation of P and Q (subsets of A*) is a certain subset of A* 113 | > 114 | > P * Q = { w \in A* | exists u, v. w = uv \and u \in P \and v \in Q } 115 | > 116 | > and the characteristic function of this set corresponds to what is given in Figure 1. This could be done in Agda with decidable languages (A* -> 2) and thus also achieve non-operational specification (values of A* -> 2 as type indices). (Assumption: a word w can be decomposed in a finite number of ways.) 117 | 118 | We wrote a comparison with set-theoretical and decidable specifications and discarded it for better focus. 119 | Yes, one can start with decidable languages and exploit the property that every finite list can be unappended in finitely many ways. 120 | Doing so would go against the principle of simple, elegant, non-operational *specifications*, since it relies on cleverness involving and motivated by decidability of language *recognition* rather than the simpler nature of language *definition*. 121 | 122 | > line 18: I think it would be helpful to later mention what properties would have required induction/coinduction? 123 | 124 | Thanks. We might refer to Abel [2016], which proves a great number of algebraic properties on tries, all of them needing coinduction/bisimulation. 125 | The key to eliminating these much more complex proofs is inheriting properties from the *indices* (languages specified as type-level predicates), which are simple and neither inductive nor coinductive. 126 | 127 | > Contribution 1, line 93: Is this meant to say that capturing the semantic essence of Brzozowski's derivatives is something new? The semantic operation was defined before the syntactic one. Brzozowski references three papers where (under different names) derivatives of sequences are used and then defines the corresponding operation on syntax. 128 | 129 | *Formalizing* (stating and proving in a formal system) the semantic essence is new, as far as we know. Generalizing this essence from languages to functions from lists might be new as well. Moreover, most recent work we know of seems to identify derivatives with regular expressions, leading readers away from the more natural and general notion in practice and its alternative realizations, such as tries. (Abel [2016] is an exception but lacks a formal semantic specification separate from the implementation itself.) 130 | 131 | See also the paragraph starting on line 717, including the mention of Brzozowski on line 720. 132 | I imagine you will agree that it is easier to make clear comparisons after the technical substance is presented (in conclusions) than before (in introduction); and yet readers want at least a general sense of comparison before the technical substance. 133 | 134 | > Contribution 2, line 99: The duality is between regular expressions and a certain subset of tries (those constructed using the language building operations). The two things named Lang in Figures 4 and 6 are not equally expressive. 135 | 136 | We agree. This comparison can be expressed more clearly, including the fact that regular expressions *and their derivatives* also naturally generalize to a much broader class of languages, as Might and Darais [2010] and follow-up work has noted (though without much specification or rigor). 137 | 138 | > line 103: Is "proof isomorphism" (throughout the paper) the same as "type isomorphism" on line 188? 139 | 140 | Yes. "Type isomorphism" is more technically specific (suiting the technical meat of the paper), while its use as "proof isomorphism" seems more suitable for the introduction and the contrast with "logical equivalence" mentioned in line 103. 141 | 142 | > Contribution 3, line 103: Is this statement considering previous work on formalisations or formal languages in general? As soon as something more general is considered instead of nullability (true/false empty word property), the lemmas (or definitions) for concatenation and closure start to look like those in Figure 2. 143 | 144 | The statement refers to work we are aware of in formal languages as well as its formalization. Have we missed something important? 145 | 146 | > line 179: What is meant exactly by "liberation"? I think it is fairly standard to consider left and right quotients (wrt. a word) of languages (sets of strings, no grammars) as monoid actions and thus reducing language membership to nullability also in a step-by-step manner. 147 | 148 | This sentence is in the context of the immediately preceding sentence, which refers to Brzozowski's paper ("Derivatives of *Regular Expressions*"---emphasis mine) and its extensions by Might and Darais [2010] etc. 149 | We are trying to encourage a shift from recent near-identification of derivatives (for parsing) as about a particular grammatical/inductive language representation back to the underlying semantic basis you mention (though also tilted toward proof relevance for parsing and other applications), so that the formal development has a simpler formal starting point (and neither inductive nor coinductive) and alternatives to grammatical representation are easier to notice. 150 | 151 | > line 283: I think it would be good to also say what is the correctness condition (in addition to the fact that it is guaranteed by types). 152 | 153 | The condition is stated on line 228 and is captured in the type of `⟦_⟧` on line 310, namely that `⟦_⟧` is both computable and converts a language into a decidable recognizer/parser. 154 | We will add a reminder at line 283. 155 | Thanks. 156 | 157 | > line 308: This left triangle looks like a map operation. Maybe a comment about why this has to be part of the syntax. 158 | 159 | The left triangle *reifies* a map operation. 160 | The reason it must be part of the inductive representation is the same as the other constructors, namely so that the core lemmas (figure 2) translate into an implementation. 161 | We will add a comment. 162 | 163 | > line 333: Maybe a comment about why f works on the left hand side of the triangle with both p and with p differentiated. 164 | 165 | Thanks. 166 | The reason is that both `p` and `δ p a` are languages/predicates. 167 | 168 | > Figures 4 and 5: there is no definition of the language (A* -> Set) denoted by an expression and thus no formal statement that it indeed denotes the language in its type. 169 | 170 | Thank you. 171 | This missing semantic function maps every `p : Lang P` to `P`, regardless of the specifics of `p` (which is why the algebraic properties transfer effortlessly from type-level languages without having to consider the structure of `p`). 172 | We agree that it's worth stating explicitly and will do so. 173 | 174 | > line 623: By a choice of an equivalence relation we are quotienting the (encapsulated) representation. It would be good to give a bit more intuition here (which quotient is this). 175 | 176 | Indeed. This comment is closely related to your remark above about the denotation of the decidable language representations and has the same answer. Equivalence of representations is *semantic* equivalence, which is to say equivalence of the indices (type-level languages) that capture the semantics. 177 | We will add this clarification. 178 | 179 | > line 723: What are transformations? 180 | 181 | The rules of differentiation, as transformations of symbolic/grammatical representations. 182 | 183 | > line 724: "based *on* the"? 184 | 185 | Fixed. Thanks! 186 | 187 | 188 | ## Review D 189 | 190 | > Section 6.2 is quite dense, and its motivation is not entirely clear to me. Is its purpose to show that existentials can be used to force the dependently-typed definitions developed throughout the paper into the simply-typed algebraic structures? If so, I believe a single paragraph would to, and would perhaps be less confusing. Furthermore, the section title is confusing: transport properties to what? We haven't defined the existentially quantifier variants yet, so it's not yet obvious that we have anywhere to transport properties to. 191 | 192 | The purpose is to show that the properties of non-decidable, type-level languages transfer effortlessly to the decidable implementations. 193 | The importance of this fact is that one gets both (a) easy proofs of useful correctness properties (without complications due to efficiency and even decidability) *and* (b) efficient implementations. 194 | 195 | This simple solution for property transfer, however, raises a "technical obstacle" (line 581): the indexed data types do not match the types of standard algebraic properties. 196 | Existential wrapping is a simple, generic solution to this obstacle. 197 | 198 | I renamed the section from "Transporting Properties" to "Transporting Properties from Specification to Implementations" 199 | 200 | > Line 86: The clause "including the use of existential quantifiers" is confusing, as it in no way makes clear what the existential quantifiers are used for. I'd suggest dropping the clause entirely. 201 | 202 | I've appended "(as in Figure 1) to that sentence". 203 | 204 | > Line 500: The types of the primitive language constructs are elided throughout the paper. Therefore, it is a bit surprising to see that the variant of star which, ostensibly, abstracts away from lists, still references the list primitives foldr and All, as well as the monoid properties. 205 | 206 | Before generalizing, the list plays two different algebraic roles: monoid and foldable. 207 | Either can be generalized independently. 208 | 209 | > Line 608: Why are the existentially-wrapped versions specifically based on the inductive variant? Is there anything specific that makes these better or more amenable? The text so far led me to believe the coinductive variant was better and more performant. 210 | 211 | It was an arbitrary choice. Both implementations (inductive and coinductive) are wrapped alike. We will say so in the text. 212 | 213 | > Line 638: Is the pattern matching in the type valid Agda syntax? Again, I don't think these liftings really add anything to the paper. 214 | 215 | Yes. All such declarations and definitions in the paper (and those elided from the paper) are fully type- (and thus correctness-)checked when the paper is typeset. 216 | 217 | * * * * * * * * * * * * * * * * * * * * 218 | 219 | From the final review: 220 |
221 | The reviewers want to thank the author for this interesting submission and are happy to accept it with the understanding that the author will make all of the following changes: 222 | 223 | 1. Clarify what's new in the context of proof mechanization or type theory, what's new in the context of parsing, and what's new in the context of formal languages & automata. In particular, some contributions do not seem to be novel in the last context. The author did not respond to Reviewer C's comment about the work of Lombardy and Sakarovitch, but their work seems to be similar to the derivative operations in the paper and the author should pay attention to it (in addition to other related works). 224 | 225 | 2. Explain the Agda syntax more for the readers who are not familiar with Agda. For example, add concrete examples showing how operators in Fig. 1 apply to concrete regular expressions and words, including the fully expanded types and values (as witnesses of the membership). The current draft has only 16 pages so space should not be an issue. 226 | 227 | 3. Clarify the difference between data and codata: both generally and in terms of difference it makes for main results of this paper. 228 | 229 | 4. Consider explaining how to recover the ordinary setting where a language is a set of words and language union is idempotent. For example, explain how the current work can be used to prove the classical theorem for the ordinary setting that the set of derivatives of a regular expression is finite modulo associativity, commutativity and idempotence of union (although the derivative defined here might require more). 230 | 231 | 5. Fix typos and other issues that have already been discussed in the reviews and the response. 232 |233 | 234 | My response: 235 | 236 | 1. I added the following paragraph to the beginning of section 7 (Related Work): 237 | 238 | > The shift from languages as sets of strings to type-level predicates (i.e., proof relevance or “parsing”) is akin to weighted automata [Schützenberger 1961; Droste and Kuske 2019] and more generally to semiring-based parsing [Chomsky and Schützenberger 1959; Goodman 1998, 1999; Liu 2004], noting that types (“Set” in Agda) form a (commutative) semiring (up to isomorphism). In particular, Lombardy and Sakarovitch [2005] investigated Brzozowski-style derivatives in this more general setting, formalizing and generalizing the work of Antimirov [1996], who decomposed derivatives into simpler components (“partial derivatives”) that sum to the full derivative and lead to efficient recognizers in the form of nondeterministic automata. 239 | 240 | Also an initial sentence to the (now) third paragraph of that section: 241 | 242 | > The main contributions of this paper are to formalization and mechanization of language recognition and parsing. 243 | 244 | 2. I added a new figure 2 with examples of languages and membership proofs, along with a new footnote 2 giving some detailed explanations. 245 | 246 | 3. I added a new footnote 10: 247 | 248 | > Inductive types (“data”) describe finitely large values and are processed recursively (as least fixed points) by pattern- matching clauses that decompose arguments. Dually, coinductive types (“codata”) describe infinitely large values and are processed corecursively (as greatest fixed points) by copattern-matching clauses that compose results. Programs on inductive types are often proved by induction, while programs on coinductive types are often proved by coinduction. (As we will see in Section 6.2, however, the simple relationship to the specification in Section 2 allows many trivial correctness proofs, needing neither induction or coinduction.) See Gordon [1995] for a tutorial on theory and techniques of coinduction. 249 | > 250 | > Tries in general represent functions, with each trie datum corresponding to a domain value. Even when each domain value is finitely large, there are often infinitely many of them (e.g., lists), so tries will naturally be infinite and thus more amenable to coinductive than inductive analysis. 251 | 252 | 4. Appended to the final paragraph of section 7: 253 | 254 | > As mentioned at the end of Section 3, one can simplify the development above somewhat by replacing proof isomorphism with the coarser notion of logical equivalence. Doing so restores union idempotence (modulo equivalence), losing proof relevance and thus re- placing parsing by recognition. One might then also prove that the set of derivatives of a regular language is finite (again, modulo equivalence). 255 | 256 | 5. Done! 257 | --------------------------------------------------------------------------------