├── Makefile ├── doc ├── strath_coat.png ├── latex │ └── thesis.pdf ├── introduction.tex │ └── type-theory.tex ├── thesis.agda-lib ├── shared │ ├── Relation │ │ └── Binary │ │ │ └── PropositionalEquality │ │ │ └── Extra.lagda │ ├── Data │ │ ├── Pred.lagda │ │ ├── Var.lagda │ │ ├── Relation.lagda │ │ └── Var │ │ │ └── Varlike.lagda │ └── Stdlib.lagda ├── type-scope-semantics.agda │ ├── Syntax │ │ ├── Type.lagda │ │ ├── MoggiML │ │ │ ├── Type.lagda │ │ │ ├── Calculus.lagda │ │ │ └── CPS.lagda │ │ ├── WeakHead.lagda │ │ ├── Normal │ │ │ └── Thinnable.lagda │ │ ├── Calculus.lagda │ │ └── Normal.lagda │ ├── Semantics │ │ ├── NormalisationByEvaluation │ │ │ ├── Specification.lagda │ │ │ ├── BetaIotaXi.lagda │ │ │ ├── BetaIota.lagda │ │ │ └── BetaIotaXiEta.lagda │ │ ├── Syntactic │ │ │ ├── Specification.lagda │ │ │ └── Instances.lagda │ │ ├── CPS │ │ │ ├── CBV.lagda │ │ │ └── CBN.lagda │ │ ├── Specification.lagda │ │ └── Printing.lagda │ ├── StateOfTheArt │ │ └── McBride05 │ │ │ ├── Base.lagda │ │ │ └── Kit.lagda │ └── Properties │ │ ├── Fusion │ │ ├── Syntactic │ │ │ ├── Specification.lagda │ │ │ └── Instances.lagda │ │ └── Specification.lagda │ │ └── Simulation │ │ └── Specification.lagda ├── robust-catch.tex ├── generic-syntax.agda │ ├── Generic │ │ ├── Syntax │ │ │ ├── LetBinder.lagda │ │ │ ├── UTLC.lagda │ │ │ ├── Extended.lagda │ │ │ ├── STLC.lagda │ │ │ ├── Bidirectional.lagda │ │ │ └── LetCounter.lagda │ │ ├── Examples │ │ │ ├── UntypedLC.lagda │ │ │ ├── HuttonsRazor.lagda │ │ │ ├── NbyE.lagda │ │ │ ├── STLC.lagda │ │ │ ├── Printing.lagda │ │ │ ├── SystemF.lagda │ │ │ ├── TypeChecking.lagda │ │ │ └── Elaboration.lagda │ │ ├── Semantics │ │ │ ├── Elaboration │ │ │ │ ├── LetBinder.lagda │ │ │ │ └── LetCounter.lagda │ │ │ ├── NbyE.lagda │ │ │ ├── Syntactic.lagda │ │ │ ├── TypeChecking.lagda │ │ │ └── Printing.lagda │ │ ├── Simulation │ │ │ └── Syntactic.lagda │ │ ├── Equality.lagda │ │ ├── Semantics.lagda │ │ ├── Scopecheck.lagda │ │ ├── Relator.lagda │ │ ├── Simulation.lagda │ │ ├── Fusion │ │ │ ├── Specialised │ │ │ │ └── Propositional.lagda │ │ │ └── Utils.lagda │ │ ├── Fusion.lagda │ │ └── Identity.lagda │ └── Motivation │ │ └── Problem │ │ └── Naive.lagda ├── generic-syntax.tex │ ├── compiling.tex │ ├── other-programs.tex │ ├── proofs.tex │ └── catalogue │ │ ├── unrolling.tex │ │ └── desugaring.tex ├── commands.tex ├── cover.tex ├── acknowledgments.tex ├── titlingpage.tex ├── thesis.tex ├── strath.tex ├── convention.tex └── type-scope-semantics.tex │ ├── normalization │ └── betaiota.tex │ └── proof-conclusion.tex ├── README.md └── .gitignore /Makefile: -------------------------------------------------------------------------------- 1 | all: 2 | make -C doc/ 3 | 4 | clean: 5 | make -C doc/ fullclean 6 | -------------------------------------------------------------------------------- /doc/strath_coat.png: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/gallais/thesis/HEAD/doc/strath_coat.png -------------------------------------------------------------------------------- /doc/latex/thesis.pdf: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/gallais/thesis/HEAD/doc/latex/thesis.pdf -------------------------------------------------------------------------------- /doc/introduction.tex/type-theory.tex: -------------------------------------------------------------------------------- 1 | \chapter{Introduction to Type Theory} 2 | \label{introduction-type-theory} 3 | -------------------------------------------------------------------------------- /README.md: -------------------------------------------------------------------------------- 1 | # Sources of Syntaxes with Binding, Their Programs, and Proofs 2 | 3 | # Compilation 4 | 5 | * Agda (2.6.1.3) 6 | * Standard Library (1.5) 7 | -------------------------------------------------------------------------------- /doc/thesis.agda-lib: -------------------------------------------------------------------------------- 1 | name: thesis 2 | depend: standard-library 3 | include: shared 4 | , introduction.agda 5 | , type-scope-semantics.agda 6 | , generic-syntax.agda 7 | -------------------------------------------------------------------------------- /.gitignore: -------------------------------------------------------------------------------- 1 | *.bak 2 | *# 3 | *~ 4 | doc/*.agda/*/*.agda 5 | doc/*.agda/*.agda 6 | doc/*.agda/*/*/*.agda 7 | *.agdai 8 | *.aux 9 | *.log 10 | *.pdf 11 | latex-agda-output/ 12 | latex/ 13 | -------------------------------------------------------------------------------- /doc/shared/Relation/Binary/PropositionalEquality/Extra.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | module Relation.Binary.PropositionalEquality.Extra where 3 | 4 | open import Relation.Binary.PropositionalEquality public 5 | 6 | cong₃ : ∀ {a b c d} {A : Set a} {B : Set b} {C : Set c} {D : Set d} 7 | (f : A → B → C → D) {x y u v s t} → x ≡ y → u ≡ v → s ≡ t → f x u s ≡ f y v t 8 | cong₃ f refl refl refl = refl 9 | 10 | \end{code} 11 | -------------------------------------------------------------------------------- /doc/type-scope-semantics.agda/Syntax/Type.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | {-# OPTIONS --safe --sized-types #-} 3 | 4 | module Syntax.Type where 5 | 6 | open import Data.Product 7 | open import Relation.Binary.PropositionalEquality 8 | 9 | infixr 20 _`→_ 10 | 11 | \end{code} 12 | %<*type> 13 | \begin{code} 14 | data Type : Set where 15 | `Unit `Bool : Type 16 | _`→_ : (σ τ : Type) → Type 17 | \end{code} 18 | % 19 | \begin{code} 20 | 21 | `→-inj : {σ₁ τ₁ σ₂ τ₂ : Type} → σ₁ `→ τ₁ ≡ σ₂ `→ τ₂ → σ₁ ≡ σ₂ × τ₁ ≡ τ₂ 22 | `→-inj refl = refl , refl 23 | \end{code} 24 | -------------------------------------------------------------------------------- /doc/type-scope-semantics.agda/Syntax/MoggiML/Type.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | module Syntax.MoggiML.Type where 3 | 4 | open import Data.Product 5 | open import Relation.Binary.PropositionalEquality 6 | 7 | infixr 20 _`→#_ #_ 8 | \end{code} 9 | %<*ctype> 10 | \begin{code} 11 | data CType : Set where 12 | `Unit : CType 13 | `Bool : CType 14 | _`→#_ : (σ τ : CType) → CType 15 | #_ : CType → CType 16 | \end{code} 17 | % 18 | \begin{code} 19 | private 20 | 21 | variable 22 | 23 | σ τ σ₁ τ₁ σ₂ τ₂ : CType 24 | 25 | `→#-inj : (σ₁ `→# τ₁) ≡ (σ₂ `→# τ₂) → σ₁ ≡ σ₂ × τ₁ ≡ τ₂ 26 | `→#-inj refl = refl , refl 27 | 28 | #-inj : # σ ≡ # τ → σ ≡ τ 29 | #-inj refl = refl 30 | \end{code} 31 | -------------------------------------------------------------------------------- /doc/robust-catch.tex: -------------------------------------------------------------------------------- 1 | \makeatletter 2 | 3 | \newrobustcmd*\OrigExecuteMetaData[2][\jobname]{% 4 | \CatchFileBetweenTags\CatchFBT@tok{#1}{#2}% 5 | \global\expandafter\CatchFBT@tok\expandafter{% 6 | \expandafter}\the\CatchFBT@tok 7 | }%\OrigExecuteMetaData 8 | 9 | \newrobustcmd*\ChkExecuteMetaData[2][\jobname]{% 10 | \CatchFileBetweenTags\CatchFBT@tok{#1}{#2}% 11 | \edef\mytokens{\detokenize\expandafter{\the\CatchFBT@tok}} 12 | \ifx\mytokens\empty\PackageError{catchfilebetweentags}{the tag #2 is not found\MessageBreak in file #1 \MessageBreak called from \jobname.tex}{use a different tag}\fi% 13 | }%\ChkExecuteMetaData 14 | 15 | \renewrobustcmd*\ExecuteMetaData[2][\jobname]{% 16 | \ChkExecuteMetaData[#1]{#2}% 17 | \OrigExecuteMetaData[#1]{#2}% 18 | } 19 | 20 | \makeatother 21 | -------------------------------------------------------------------------------- /doc/generic-syntax.agda/Generic/Syntax/LetBinder.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | {-# OPTIONS --safe --sized-types #-} 3 | 4 | module Generic.Syntax.LetBinder where 5 | 6 | open import Size 7 | open import Data.Bool 8 | open import Data.Product 9 | open import Agda.Builtin.List 10 | open import Agda.Builtin.Equality 11 | open import Function 12 | open import Relation.Unary 13 | 14 | open import Generic.Syntax 15 | 16 | private 17 | variable 18 | I : Set 19 | σ τ : I 20 | d : Desc I 21 | s : Size 22 | 23 | 24 | module _ {I : Set} where 25 | 26 | \end{code} 27 | %<*letcode> 28 | \begin{code} 29 | Let : Desc I 30 | Let = `σ (I × I) $ λ (σ , τ) → 31 | `X [] σ (`X (σ ∷ []) τ (`∎ τ)) 32 | \end{code} 33 | % 34 | %<*letpattern> 35 | \begin{code} 36 | pattern `let'_`in'_ e t = (_ , e , t , refl) 37 | pattern `let_`in_ e t = `con (`let' e `in' t) 38 | \end{code} 39 | % 40 | 41 | \begin{code} 42 | embed : ∀[ Tm d s σ ⇒ Tm (d `+ Let) s σ ] 43 | embed = map^Tm (MkDescMorphism (true ,_)) 44 | \end{code} 45 | -------------------------------------------------------------------------------- /doc/generic-syntax.agda/Generic/Syntax/UTLC.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | {-# OPTIONS --safe --sized-types #-} 3 | 4 | module Generic.Syntax.UTLC where 5 | 6 | open import Size 7 | open import Data.Unit 8 | open import Data.Bool 9 | open import Data.Product 10 | open import Data.List.Base hiding ([_]) 11 | open import Function 12 | open import Relation.Binary.PropositionalEquality hiding ([_]) 13 | 14 | open import Data.Var 15 | open import Generic.Syntax 16 | 17 | \end{code} 18 | %<*ULC> 19 | \begin{code} 20 | UTLC : Desc ⊤ 21 | UTLC = `σ Bool $ λ isApp → if isApp 22 | then `X [] tt (`X [] tt (`∎ tt)) 23 | else `X (tt ∷ []) tt (`∎ tt) 24 | \end{code} 25 | % 26 | \begin{code} 27 | private 28 | module DISPLAYONLY where 29 | \end{code} 30 | %<*LCpat> 31 | \begin{code} 32 | pattern `app f t = `con (true , f , t , refl) 33 | pattern `lam b = `con (false , b , refl) 34 | \end{code} 35 | % 36 | \begin{code} 37 | pattern `app' f t = (true , f , t , refl) 38 | pattern `lam' b = (false , b , refl) 39 | pattern `app f t = `con (`app' f t) 40 | pattern `lam b = `con (`lam' b) 41 | \end{code} 42 | %<*LCid> 43 | \begin{code} 44 | id^U : Tm UTLC ∞ tt [] 45 | id^U = `lam (`var z) 46 | \end{code} 47 | % 48 | -------------------------------------------------------------------------------- /doc/shared/Data/Pred.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | {-# OPTIONS --safe --sized-types #-} 3 | 4 | module Data.Pred where 5 | 6 | open import Data.Var hiding (_<$>_) 7 | open import Data.Environment 8 | open import Agda.Builtin.List 9 | open import Function 10 | open import Relation.Unary using (IUniversal; _⇒_) 11 | 12 | private 13 | variable 14 | I : Set 15 | σ : I 16 | Γ Δ Θ : List I 17 | T : I ─Scoped 18 | 19 | record Pred (T : I ─Scoped) : Set₁ where 20 | constructor mkPred 21 | field pred : ∀ σ → ∀[ T σ ⇒ const Set ] 22 | open Pred public 23 | 24 | record All (𝓟 : Pred T) (Γ : List I) (ρ : (Γ ─Env) T Δ) : Set where 25 | constructor packᴾ 26 | field lookupᴾ : (k : Var σ Γ) → pred 𝓟 σ (lookup ρ k) 27 | open All public 28 | 29 | private 30 | variable 31 | 𝓟 : Pred T 32 | 33 | _<$>ᴾ_ : {f : {i : I} → T i Δ → T i Θ} → 34 | (∀ {i} {t : T i Δ} → pred 𝓟 i t → pred 𝓟 i (f t)) → 35 | {ρ : (Γ ─Env) T Δ} → All 𝓟 Γ ρ → All 𝓟 Γ (f <$> ρ) 36 | lookupᴾ (F <$>ᴾ ρ) k = F (lookupᴾ ρ k) 37 | 38 | εᴾ : All 𝓟 [] (([] ─Env) T Δ ∋ ε) 39 | lookupᴾ εᴾ () 40 | 41 | infixl 20 _∙ᴾ_ 42 | _∙ᴾ_ : ∀ {ρ : (Γ ─Env) T Δ} {v : T σ Δ} → All 𝓟 Γ ρ → pred 𝓟 σ v → All 𝓟 (σ ∷ Γ) (ρ ∙ v) 43 | lookupᴾ (ρ ∙ᴾ v) z = v 44 | lookupᴾ (ρ ∙ᴾ v) (s k) = lookupᴾ ρ k 45 | \end{code} 46 | -------------------------------------------------------------------------------- /doc/generic-syntax.agda/Generic/Examples/UntypedLC.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | module Generic.Examples.UntypedLC where 3 | 4 | open import Size 5 | open import Data.Unit 6 | open import Data.Bool 7 | open import Data.Product 8 | open import Data.List.Base hiding ([_]) 9 | open import Function 10 | open import Relation.Binary.PropositionalEquality hiding ([_]) 11 | 12 | open import Relation.Unary 13 | open import Data.Var 14 | open import Generic.Syntax 15 | 16 | -- typesetting only 17 | module TYPESETTING-UTLC where 18 | \end{code} 19 | %<*dataulc> 20 | \begin{code} 21 | data UTLC : ⊤ ─Scoped where 22 | `var : ∀[ Var _ ⇒ UTLC _ ] 23 | -- choice of two constructors: 24 | `app : ∀[ UTLC _ ⇒ UTLC _ ⇒ UTLC _ ] 25 | `lam : ∀[ (tt ∷_) ⊢ UTLC _ ⇒ UTLC _ ] 26 | \end{code} 27 | % 28 | 29 | \end{code} 30 | %<*ulc> 31 | \begin{code} 32 | UTLC : Desc ⊤ 33 | -- var will be freely adjoined 34 | UTLC = `σ Bool $ λ isApp → if isApp 35 | then `X [] tt (`X [] tt (`∎ tt)) 36 | else `X (tt ∷ []) tt (`∎ tt) 37 | \end{code} 38 | % 39 | 40 | %<*patterns> 41 | \begin{code} 42 | pattern `app f t = `con (true , f , t , refl) 43 | pattern `lam b = `con (false , b , refl) 44 | \end{code} 45 | % 46 | %<*identity> 47 | \begin{code} 48 | id^U : TM UTLC tt 49 | id^U = `lam (`var z) 50 | \end{code} 51 | % 52 | -------------------------------------------------------------------------------- /doc/type-scope-semantics.agda/Semantics/NormalisationByEvaluation/Specification.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | 3 | module Semantics.NormalisationByEvaluation.Specification where 4 | 5 | open import Semantics.Specification 6 | open import Data.List.Base 7 | open import Data.Var 8 | open import Data.Environment 9 | open import Syntax.Type 10 | open import Syntax.Calculus 11 | open import Relation.Unary 12 | open import Function.Base 13 | 14 | private 15 | variable 16 | σ : Type 17 | Γ Δ : List Type 18 | 19 | \end{code} 20 | %<*recnbe> 21 | \begin{code} 22 | record NBE (Model : Type ─Scoped) (Nf : Type ─Scoped) : Set₁ where 23 | field Sem : Semantics Model Model 24 | embed : ∀[ Var σ ⇒ Model σ ] 25 | reify : ∀[ Model σ ⇒ Nf σ ] 26 | \end{code} 27 | % 28 | \begin{code} 29 | 30 | \end{code} 31 | %<*eval> 32 | \begin{code} 33 | eval : ∀[ Term σ ⇒ Model σ ] 34 | eval = semantics Sem (pack embed) 35 | \end{code} 36 | % 37 | \begin{code} 38 | 39 | \end{code} 40 | %<*nbe> 41 | \begin{code} 42 | nbe : ∀[ Term σ ⇒ Nf σ ] 43 | nbe = reify ∘ eval 44 | \end{code} 45 | % 46 | %<*test> 47 | \begin{code} 48 | test : Nf (`Bool `→ `Unit `→ `Unit) [] 49 | test = nbe (`lam (`lam (`app (`lam (`var z)) 50 | (`ifte (`var (s z)) `one (`var z))))) 51 | \end{code} 52 | % 53 | -------------------------------------------------------------------------------- /doc/generic-syntax.tex/compiling.tex: -------------------------------------------------------------------------------- 1 | \chapter{Compiler Passes as Semantics} 2 | 3 | In the previous chapter we have seen various generic semantics one may be 4 | interested in when working on a deeply embedded language: renaming, 5 | substitution, printing with names or evaluation. All of these fit neatly 6 | in the \AR{Semantics} framework. 7 | 8 | Now we wish to focus on the kind of traversals one may find in a compiler 9 | pipeline such as a scopechecker, an elaboration function from an untyped 10 | surface syntax to an intrinsically typed syntax, or an optimisation pass 11 | inlining definitions used at most once. 12 | 13 | All of the examples but the scopechecker are instances of \AR{Semantics}. 14 | Some, like the scoping (\cref{section:genericscoping}), desugaring 15 | (\cref{section:letbinding}), and inlining (\cref{section:inlining}) passes 16 | will be fully generic while others like the typechecking 17 | (\cref{section:typechecking}) and elaboration (\cref{section:elaboration}) 18 | passes will correspond to a specific language and its particular type 19 | system. 20 | 21 | \input{generic-syntax.tex/catalogue/scopechecking.tex} 22 | \input{generic-syntax.tex/catalogue/typechecking.tex} 23 | \input{generic-syntax.tex/catalogue/elaborating.tex} 24 | \input{generic-syntax.tex/catalogue/desugaring.tex} 25 | \input{generic-syntax.tex/catalogue/inlining.tex} 26 | -------------------------------------------------------------------------------- /doc/type-scope-semantics.agda/StateOfTheArt/McBride05/Base.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | module StateOfTheArt.McBride05.Base where 3 | 4 | open import Data.List hiding ([_]; lookup) 5 | open import Data.Var hiding (_<$>_) 6 | open import Data.Environment 7 | open import Relation.Unary 8 | open import Function 9 | 10 | data Type : Set where 11 | base : Type 12 | arr : Type → Type → Type 13 | 14 | private 15 | variable 16 | σ τ : Type 17 | Γ Δ : List Type 18 | \end{code} 19 | 20 | %<*stlc> 21 | \begin{code} 22 | data Tm : Type ─Scoped where 23 | `var : ∀[ Var σ ⇒ Tm σ ] 24 | `app : ∀[ Tm (arr σ τ) ⇒ Tm σ ⇒ Tm τ ] 25 | `lam : ∀[ (σ ∷_) ⊢ Tm τ ⇒ Tm (arr σ τ) ] 26 | \end{code} 27 | % 28 | 29 | \begin{code} 30 | 31 | module RenSub where 32 | \end{code} 33 | %<*ren> 34 | \begin{code} 35 | ren : (Γ ─Env) Var Δ → Tm σ Γ → Tm σ Δ 36 | ren ρ (`var v) = `var (lookup ρ v) 37 | ren ρ (`app f t) = `app (ren ρ f) (ren ρ t) 38 | ren ρ (`lam b) = `lam (ren ρ′ b) 39 | where ρ′ = (s <$> ρ) ∙ z 40 | \end{code} 41 | % 42 | 43 | %<*sub> 44 | \begin{code} 45 | sub : (Γ ─Env) Tm Δ → Tm σ Γ → Tm σ Δ 46 | sub ρ (`var v) = lookup ρ v 47 | sub ρ (`app f t) = `app (sub ρ f) (sub ρ t) 48 | sub ρ (`lam b) = `lam (sub ρ′ b) 49 | where ρ′ = (ren (pack s) <$> ρ) ∙ `var z 50 | \end{code} 51 | % 52 | -------------------------------------------------------------------------------- /doc/generic-syntax.agda/Generic/Syntax/Extended.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | module Generic.Syntax.Extended where 3 | 4 | open import Generic.Syntax 5 | open import Data.Var 6 | open import Data.List.Base 7 | open import Syntax.Type 8 | open import Function 9 | 10 | private 11 | variable 12 | σ τ : Type 13 | 14 | \end{code} 15 | %<*tag:decl> 16 | \begin{code} 17 | data `Term : Set where 18 | \end{code} 19 | % 20 | %<*tag:lam> 21 | \begin{code} 22 | Lam : Type → Type → `Term 23 | \end{code} 24 | % 25 | %<*tag:base> 26 | \begin{code} 27 | One TT FF : `Term 28 | \end{code} 29 | % 30 | %<*tag:struct> 31 | \begin{code} 32 | App : Type → Type → `Term 33 | Ifte : Type → `Term 34 | \end{code} 35 | % 36 | \begin{code} 37 | 38 | \end{code} 39 | %<*term:decl> 40 | \begin{code} 41 | Term : Desc Type 42 | Term = `σ `Term $ λ where 43 | \end{code} 44 | % 45 | %<*term:lam> 46 | \begin{code} 47 | (Lam σ τ) → `X (σ ∷ []) τ (`∎ (σ `→ τ)) 48 | \end{code} 49 | % 50 | %<*term:base> 51 | \begin{code} 52 | One → `∎ `Unit 53 | TT → `∎ `Bool 54 | FF → `∎ `Bool 55 | \end{code} 56 | % 57 | %<*term:struct> 58 | \begin{code} 59 | (App σ τ) → `X [] (σ `→ τ) (`X [] σ (`∎ τ)) 60 | (Ifte σ) → `X [] `Bool (`X [] σ (`X [] σ (`∎ σ))) 61 | \end{code} 62 | % 63 | \begin{code} 64 | -------------------------------------------------------------------------------- /doc/generic-syntax.agda/Generic/Syntax/STLC.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | {-# OPTIONS --safe --sized-types #-} 3 | 4 | module Generic.Syntax.STLC where 5 | 6 | open import Size 7 | open import Data.Bool 8 | open import Data.Product 9 | open import Agda.Builtin.Equality 10 | open import Agda.Builtin.List 11 | open import Generic.Syntax 12 | open import Data.Var 13 | open import Generic.Syntax.Bidirectional using (Type ; α ; _`→_) 14 | open import Function 15 | 16 | private 17 | variable 18 | σ : Type 19 | 20 | \end{code} 21 | %<*stlc> 22 | \begin{AgdaAlign} 23 | %<*tag> 24 | \begin{code} 25 | data `STLC : Set where 26 | App Lam : Type → Type → `STLC 27 | \end{code} 28 | % 29 | \begin{AgdaSuppressSpace} 30 | %<*desc> 31 | \begin{code} 32 | STLC : Desc Type 33 | STLC = `σ `STLC $ λ where 34 | (App σ τ) → `X [] (σ `→ τ) (`X [] σ (`∎ τ)) 35 | (Lam σ τ) → `X (σ ∷ []) τ (`∎ (σ `→ τ)) 36 | \end{code} 37 | % 38 | \end{AgdaSuppressSpace} 39 | \end{AgdaAlign} 40 | % 41 | %<*patST> 42 | \begin{code} 43 | pattern `app f t = `con (App _ _ , f , t , refl) 44 | pattern `lam b = `con (Lam _ _ , b , refl) 45 | \end{code} 46 | % 47 | \begin{code} 48 | _ : Tm STLC ∞ ((α `→ α) `→ (α `→ α)) [] 49 | _ = `lam (`lam (`app (`var (s z)) (`var z))) 50 | \end{code} 51 | 52 | %<*STid> 53 | \begin{code} 54 | id^S : Tm STLC ∞ (σ `→ σ) [] 55 | id^S = `lam (`var z) 56 | \end{code} 57 | % 58 | -------------------------------------------------------------------------------- /doc/type-scope-semantics.agda/Syntax/MoggiML/Calculus.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | module Syntax.MoggiML.Calculus where 3 | 4 | open import Syntax.MoggiML.Type 5 | open import Data.List.Base using (List; []; _∷_) 6 | open import Data.Var 7 | open import Data.Environment 8 | open import Relation.Unary 9 | 10 | private 11 | 12 | variable 13 | 14 | σ τ : CType 15 | 16 | \end{code} 17 | %<*calculus> 18 | \begin{code} 19 | data ML : CType ─Scoped where 20 | `var : ∀[ Var σ ⇒ ML σ ] 21 | `app : ∀[ ML (σ `→# τ) ⇒ ML σ ⇒ ML (# τ) ] 22 | `lam : ∀[ (σ ∷_) ⊢ ML (# τ) ⇒ ML (σ `→# τ) ] 23 | `one : ∀[ ML `Unit ] 24 | `tt `ff : ∀[ ML `Bool ] 25 | `ifte : ∀[ ML `Bool ⇒ ML (# σ) ⇒ ML (# σ) ⇒ ML (# σ) ] 26 | `ret : ∀[ ML σ ⇒ ML (# σ) ] 27 | `bind : ∀[ ML (# σ) ⇒ ML (σ `→# τ) ⇒ ML (# τ) ] 28 | \end{code} 29 | % 30 | %<*thcalculus> 31 | \begin{code} 32 | th^ML : Thinnable (ML σ) 33 | th^ML (`var v) ρ = `var (lookup ρ v) 34 | th^ML (`app f t) ρ = `app (th^ML f ρ) (th^ML t ρ) 35 | th^ML (`lam b) ρ = `lam (th^ML b (th^Env th^Var ρ extend ∙ z)) 36 | th^ML `one ρ = `one 37 | th^ML `tt ρ = `tt 38 | th^ML `ff ρ = `ff 39 | th^ML (`ifte t l r) ρ = `ifte (th^ML t ρ) (th^ML l ρ) (th^ML r ρ) 40 | th^ML (`ret t) ρ = `ret (th^ML t ρ) 41 | th^ML (`bind m f) ρ = `bind (th^ML m ρ) (th^ML f ρ) 42 | \end{code} 43 | % 44 | -------------------------------------------------------------------------------- /doc/commands.tex: -------------------------------------------------------------------------------- 1 | %%%%%%%%%% AGDA ALIASES 2 | 3 | \newcommand{\APT}{\AgdaPrimitiveType} 4 | \newcommand{\AK}{\AgdaKeyword} 5 | \newcommand{\AM}{\AgdaModule} 6 | \newcommand{\AS}{\AgdaSymbol} 7 | \newcommand{\AStr}{\AgdaString} 8 | \newcommand{\AN}{\AgdaNumber} 9 | \newcommand{\AD}{\AgdaDatatype} 10 | \newcommand{\AF}{\AgdaFunction} 11 | \newcommand{\AR}{\AgdaRecord} 12 | \newcommand{\ARF}{\AgdaField} 13 | \newcommand{\AB}{\AgdaBound} 14 | \newcommand{\AIC}{\AgdaInductiveConstructor} 15 | 16 | \newcommand{\Set}{{\ensuremath{\mathbf{Set}}}} 17 | 18 | % common notions 19 | 20 | \newcommand{\scoped}[1]{(#1 \AR{−Scoped})} 21 | \newcommand{\pholder}{\ensuremath{\cdot{}}} 22 | \newcommand{\scopeandtypesafe}{scope- and type-safe} 23 | 24 | % language 25 | 26 | \newcommand{\nonterminal}[1]{\ensuremath{⟨#1⟩}} 27 | \newcommand{\type}{\mathtt{Type}} 28 | \newcommand{\unit}{\ensuremath{\mathtt{Unit}}} 29 | \newcommand{\bool}{\ensuremath{\mathtt{Bool}}} 30 | \newcommand{\arrow}[2]{\ensuremath{#1}~{→}~\ensuremath{#2}} 31 | 32 | \newcommand{\term}{\mathtt{Term}} 33 | \newcommand{\app}[2]{\ensuremath{#1\,#2}} 34 | \newcommand{\lam}[2]{\ensuremath{λ#1.\,#2}} 35 | \newcommand{\uni}{\ensuremath{()}} 36 | \newcommand{\tru}{\ensuremath{\mathtt{true}}} 37 | \newcommand{\fls}{\ensuremath{\mathtt{false}}} 38 | \newcommand{\ifte}[3]{\ensuremath{\mathtt{if}\,#1\,\mathtt{then}\,#2\,\mathtt{else}\,#3}} 39 | 40 | 41 | \newcommand{\semrec}{\AR{Semantics}} 42 | \newcommand{\semfun}{\AF{semantics}} 43 | -------------------------------------------------------------------------------- /doc/type-scope-semantics.agda/Semantics/Syntactic/Specification.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | module Semantics.Syntactic.Specification where 3 | 4 | open import Data.List.Base using (List; []; _∷_) 5 | open import Syntax.Type 6 | open import Syntax.Calculus 7 | open import Data.Var 8 | open import Data.Environment 9 | open import Semantics.Specification as Semantics 10 | open import Relation.Unary 11 | 12 | private 13 | variable 14 | σ : Type 15 | Γ : List Type 16 | 𝓣 : Type ─Scoped 17 | \end{code} 18 | %<*syntactic> 19 | \begin{code} 20 | record Syntactic (𝓣 : Type ─Scoped) : Set where 21 | field zro : ∀[ (σ ∷_) ⊢ 𝓣 σ ] 22 | th^𝓣 : Thinnable (𝓣 σ) 23 | var : ∀[ 𝓣 σ ⇒ Term σ ] 24 | \end{code} 25 | % 26 | \begin{code} 27 | 28 | -- Using copatterns in the definition of syntactic guarantees that these things are 29 | -- not unfolded when normalising goals thus making them more readable. 30 | 31 | \end{code} 32 | %<*syntacticsem> 33 | \begin{code} 34 | module _ (𝓢 : Syntactic 𝓣) where 35 | 36 | open Syntactic 𝓢 37 | 38 | syntactic : Semantics 𝓣 Term 39 | Semantics.th^𝓥 syntactic = th^𝓣 40 | Semantics.var syntactic = var 41 | Semantics.lam syntactic = λ b → `lam (b extend zro) 42 | Semantics.app syntactic = `app 43 | Semantics.one syntactic = `one 44 | Semantics.tt syntactic = `tt 45 | Semantics.ff syntactic = `ff 46 | Semantics.ifte syntactic = `ifte 47 | \end{code} 48 | % 49 | -------------------------------------------------------------------------------- /doc/generic-syntax.agda/Generic/Examples/HuttonsRazor.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | {-# OPTIONS --safe --sized-types #-} 3 | 4 | module Generic.Examples.HuttonsRazor where 5 | 6 | open import Size 7 | open import Data.Empty 8 | open import Data.Unit 9 | open import Data.Bool 10 | open import Data.Product 11 | open import Data.Nat 12 | open import Data.List 13 | open import Relation.Binary.PropositionalEquality 14 | 15 | open import Data.Environment 16 | open import Generic.Syntax 17 | open import Generic.Semantics 18 | 19 | -- Hutton's razor as a minimalistic example of a language 20 | -- one may want to evaluate 21 | 22 | HuttRaz : Desc ⊤ 23 | HuttRaz = `σ ℕ (λ _ → `∎ tt) 24 | `+ `X [] tt (`X [] tt (`∎ tt)) 25 | 26 | infixr 5 _[+]_ 27 | pattern lit n = `con (true , n , refl) 28 | pattern _[+]_ e f = `con (false , e , f , refl) 29 | 30 | -- Because there are no variables whatsoever in this simple 31 | -- language we can simply associated values of the empty to 32 | -- them. The computation itself will deliver a natural number. 33 | 34 | Eval : Semantics HuttRaz (λ _ _ → ⊥) (λ _ _ → ℕ) 35 | Semantics.th^𝓥 Eval = ⊥-elim 36 | Semantics.var Eval = ⊥-elim 37 | Semantics.alg Eval = case proj₁ (λ { (m , n , _) → m + n }) 38 | 39 | eval : Tm HuttRaz ∞ tt [] → ℕ 40 | eval = Semantics.closed Eval 41 | 42 | -- And, sure enough, we are able to run these expressions 43 | 44 | 3+2 : eval (lit 3 [+] lit 2) ≡ 5 45 | 3+2 = refl 46 | 47 | [2+6]+0 : eval ((lit 2 [+] lit 6) [+] lit 0) ≡ 8 48 | [2+6]+0 = refl 49 | \end{code} 50 | -------------------------------------------------------------------------------- /doc/cover.tex: -------------------------------------------------------------------------------- 1 | \pagestyle{empty} 2 | 3 | \savegeometry{maingeometry} 4 | \newgeometry{margin=0pt, bottom=5em} 5 | 6 | \definecolor{myyellow}{HTML}{fef1c4} 7 | \definecolor{myred}{HTML}{ed2a30} 8 | \definecolor{myblack}{HTML}{131313} 9 | 10 | \pagecolor{myyellow}\afterpage{\nopagecolor} 11 | 12 | \noindent\begin{minipage}{\paperwidth} 13 | \fontsize{40}{48}\selectfont 14 | \vspace*{2em} 15 | \begin{center} 16 | \textsc{\textbf{Syntaxes with Binding,\\ 17 | Their Programs, and Proofs}} 18 | \end{center} 19 | \end{minipage} 20 | 21 | \begin{vplace}[1] 22 | \renewcommand{\baselinestretch}{1} 23 | 24 | \noindent\colorbox{myblack}{\begin{minipage}{\paperwidth}~\end{minipage}} 25 | \noindent\colorbox{myyellow}{\begin{minipage}{\paperwidth}~\end{minipage}} 26 | \noindent\colorbox{myred}{\begin{minipage}{\paperwidth}% 27 | \fontsize{50}{60}\selectfont 28 | \boldmath 29 | \begin{mathpar} 30 | \color{myyellow}{\inferrule{\Box~(\mathcal{V}~\sigma \Rightarrow \mathcal{C}~\tau)} 31 | {\mathcal{C}~(\sigma \rightarrow \tau)}} 32 | \end{mathpar} 33 | \unboldmath 34 | \end{minipage}} 35 | \noindent\colorbox{myyellow}{\begin{minipage}{\paperwidth}~\end{minipage}} 36 | \noindent\colorbox{myblack}{\begin{minipage}{\paperwidth}~\end{minipage}} 37 | \end{vplace} 38 | 39 | 40 | \noindent\fcolorbox{myred}{myred}{ 41 | \begin{minipage}{\paperwidth} 42 | \vspace{8pt} 43 | \fontsize{25}{30}\selectfont 44 | \hspace{8pt}\color{myyellow}{\textsc{\textbf{Allais Guillaume}}} 45 | \vspace{8pt} 46 | \end{minipage}} 47 | 48 | \loadgeometry{maingeometry} 49 | -------------------------------------------------------------------------------- /doc/generic-syntax.agda/Generic/Examples/NbyE.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | {-# OPTIONS --sized-types #-} 3 | 4 | module Generic.Examples.NbyE where 5 | 6 | open import Size 7 | open import Data.Bool.Base 8 | open import Data.List.Base 9 | open import Data.Maybe 10 | open import Data.Product 11 | open import Data.Unit 12 | open import Function 13 | open import Relation.Unary 14 | open import Relation.Binary.PropositionalEquality 15 | 16 | open import Data.Var using (_─Scoped) 17 | open import Data.Var.Varlike 18 | open import Data.Environment hiding (_$$_) 19 | open import Generic.Syntax 20 | open import Generic.Syntax.UTLC 21 | open import Generic.Semantics.NbyE 22 | 23 | private 24 | variable 25 | I : Set 26 | 𝓥 𝓒 : I ─Scoped 27 | σ τ : I 28 | 29 | 30 | \end{code} 31 | %<*nbepatterns> 32 | \begin{code} 33 | pattern LAM f = C (false , f , refl) 34 | pattern APP' f t = (true , f , t , refl) 35 | \end{code} 36 | % 37 | 38 | \begin{code} 39 | \end{code} 40 | %<*app> 41 | \begin{code} 42 | _$$_ : ∀[ Kripke 𝓥 𝓒 (σ ∷ []) τ ⇒ (𝓥 σ ⇒ 𝓒 τ) ] 43 | f $$ t = extract f (ε ∙ t) 44 | \end{code} 45 | % 46 | 47 | %<*nbelc> 48 | \begin{code} 49 | norm^LC : ∀[ Tm UTLC ∞ tt ⇒ Maybe ∘ Tm UTLC ∞ tt ] 50 | norm^LC = norm $ λ where 51 | (APP' (LAM f) t) → f $$ t -- redex 52 | t → C t -- value 53 | \end{code} 54 | % 55 | \begin{code} 56 | open import Relation.Binary.PropositionalEquality hiding ([_] ; refl) 57 | 58 | \end{code} 59 | %<*example> 60 | \begin{code} 61 | _ : norm^LC (`app id^U (`app id^U id^U)) ≡ just id^U 62 | _ = refl 63 | \end{code} 64 | % 65 | -------------------------------------------------------------------------------- /doc/generic-syntax.agda/Generic/Semantics/Elaboration/LetBinder.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | {-# OPTIONS --safe --sized-types #-} 3 | 4 | module Generic.Semantics.Elaboration.LetBinder where 5 | 6 | open import Size 7 | open import Data.Product 8 | open import Agda.Builtin.List 9 | open import Function 10 | open import Relation.Unary 11 | 12 | open import Data.Environment 13 | open import Generic.Syntax 14 | open import Generic.Syntax.LetBinder 15 | open import Generic.Semantics 16 | open import Generic.Semantics.Syntactic 17 | 18 | private 19 | variable 20 | I : Set 21 | d : Desc I 22 | σ : I 23 | Γ Δ : List I 24 | s : Size 25 | 26 | -- Elaborating away a single let-binder. The algebra is defined by case analysis 27 | -- over the constructors: 28 | 29 | -- * let-binders are inlined thanks to the substitution _[_/0] which takes two 30 | -- arguments t and e and instatiates the first free variable of t with e. 31 | 32 | -- * the other constructors are left as is by reusing Substitution's algebra 33 | 34 | 35 | \end{code} 36 | %<*unletcode> 37 | \begin{code} 38 | UnLet : Semantics (d `+ Let) (Tm d ∞) (Tm d ∞) 39 | Semantics.th^𝓥 UnLet = th^Tm 40 | Semantics.var UnLet = id 41 | Semantics.alg UnLet = case (Semantics.alg Sub) $ λ where 42 | (`let' e `in' t) → extract t (ε ∙ e) 43 | \end{code} 44 | % 45 | \begin{code} 46 | 47 | unLet : (Γ ─Env) (Tm d ∞) Δ → Tm (d `+ Let) s σ Γ → Tm d ∞ σ Δ 48 | unLet ρ t = Semantics.semantics UnLet ρ t 49 | 50 | \end{code} 51 | %<*unlet> 52 | \begin{code} 53 | unlet : ∀[ Tm (d `+ Let) ∞ σ ⇒ Tm d ∞ σ ] 54 | unlet = Semantics.semantics UnLet id^Tm 55 | \end{code} 56 | % 57 | -------------------------------------------------------------------------------- /doc/generic-syntax.agda/Generic/Examples/STLC.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | module Generic.Examples.STLC where 3 | 4 | open import Size 5 | open import Data.Bool 6 | open import Data.Product 7 | open import Agda.Builtin.Equality 8 | open import Agda.Builtin.List 9 | open import Generic.Syntax 10 | open import Data.Var 11 | open import Generic.Examples.TypeChecking using (Type); open Type 12 | open import Function 13 | open import Relation.Unary 14 | 15 | private 16 | variable 17 | σ τ : Type 18 | 19 | 20 | -- typesetting only 21 | module TYPESETTING-STLC where 22 | \end{code} 23 | %<*datastlc> 24 | \begin{code} 25 | data STLC : Type ─Scoped where 26 | `var : ∀[ Var σ ⇒ STLC σ ] 27 | -- choice of two constructors: 28 | `app : ∀[ STLC (σ `→ τ) ⇒ 29 | STLC σ ⇒ STLC τ ] 30 | `lam : ∀[ (σ ∷_) ⊢ STLC τ ⇒ STLC (σ `→ τ) ] 31 | \end{code} 32 | % 33 | 34 | \end{code} 35 | %<*tag> 36 | \begin{code} 37 | data `STLC : Set where 38 | App Lam : Type → Type → `STLC 39 | \end{code} 40 | % 41 | 42 | %<*stlc> 43 | \begin{code} 44 | STLC : Desc Type 45 | -- var will be freely adjoined 46 | STLC = `σ `STLC $ λ where 47 | (App σ τ) → `X [] (σ `→ τ) 48 | (`X [] σ (`∎ τ)) 49 | (Lam σ τ) → `X (σ ∷ []) τ (`∎ (σ `→ τ)) 50 | \end{code} 51 | % 52 | \begin{code} 53 | 54 | \end{code} 55 | %<*patterns> 56 | \begin{code} 57 | pattern `app f t = `con (App _ _ , f , t , refl) 58 | pattern `lam b = `con (Lam _ _ , b , refl) 59 | \end{code} 60 | % 61 | \begin{code} 62 | 63 | \end{code} 64 | %<*identity> 65 | \begin{code} 66 | id^S : TM STLC (σ `→ σ) 67 | id^S = `lam (`var z) 68 | \end{code} 69 | % 70 | -------------------------------------------------------------------------------- /doc/type-scope-semantics.agda/Syntax/MoggiML/CPS.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | module Syntax.MoggiML.CPS where 3 | 4 | open import Syntax.Type 5 | open import Syntax.Calculus 6 | open import Semantics.Syntactic.Instances 7 | open import Syntax.MoggiML.Type 8 | open import Syntax.MoggiML.Calculus 9 | 10 | open import Data.List.Base 11 | open import Data.Var 12 | open import Data.Environment hiding (_<$>_) 13 | open import Function renaming (_$′_ to _$_; _∘′_ to _∘_) using () 14 | 15 | private 16 | 17 | variable 18 | 19 | r : Type 20 | σ τ : CType 21 | Γ : List CType 22 | \end{code} 23 | %<*cps> 24 | \begin{code} 25 | mutual 26 | 27 | #CPS[_] : Type → CType → Type 28 | #CPS[ r ] σ = (CPS[ r ] σ `→ r) `→ r 29 | 30 | CPS[_] : Type → CType → Type 31 | CPS[ r ] `Bool = `Bool 32 | CPS[ r ] `Unit = `Unit 33 | CPS[ r ] (σ `→# τ) = CPS[ r ] σ `→ #CPS[ r ] τ 34 | CPS[ r ] (# σ) = #CPS[ r ] σ 35 | \end{code} 36 | % 37 | %<*elabtype> 38 | \begin{code} 39 | cps : ML σ Γ → Term (CPS[ r ] σ) (map CPS[ r ] Γ) 40 | \end{code} 41 | % 42 | %<*elab> 43 | \begin{code} 44 | cps (`var v) = `var (CPS[ _ ] <$> v) 45 | cps (`app f t) = `app (cps f) (cps t) 46 | cps (`lam b) = `lam (cps b) 47 | cps `one = `one 48 | cps `tt = `tt 49 | cps `ff = `ff 50 | cps (`ifte b l r) = `ifte (cps b) (cps l) (cps r) 51 | \end{code} 52 | % 53 | %<*elabret> 54 | \begin{code} 55 | cps (`ret t) = `lam (`app (`var z) (th^Term (cps t) extend)) 56 | \end{code} 57 | % 58 | %<*elabbind> 59 | \begin{code} 60 | cps (`bind m f) = `lam $ `app m' $ `lam $ `app (`app f' (`var z)) (`var (s z)) 61 | where m' = th^Term (cps m) (pack s) 62 | f' = th^Term (cps f) (pack (s ∘ s)) 63 | \end{code} 64 | % 65 | -------------------------------------------------------------------------------- /doc/shared/Stdlib.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | module Stdlib where 3 | 4 | open import Data.Product using (_×_) 5 | open import Data.List.Base 6 | 7 | private 8 | variable 9 | A B : Set 10 | \end{code} 11 | 12 | %<*constant> 13 | \begin{code} 14 | const : Set → (A → Set) 15 | const P x = P 16 | \end{code} 17 | % 18 | 19 | %<*product> 20 | \begin{code} 21 | _∙×_ : (P Q : A → Set) → (A → Set) 22 | (P ∙× Q) x = P x × Q x 23 | \end{code} 24 | % 25 | 26 | %<*arrow> 27 | \begin{code} 28 | _⇒_ : (P Q : A → Set) → (A → Set) 29 | (P ⇒ Q) x = P x → Q x 30 | \end{code} 31 | % 32 | %<*forall> 33 | \begin{code} 34 | ∀[_] : (A → Set) → Set 35 | ∀[_] P = ∀ {x} → P x 36 | \end{code} 37 | % 38 | %<*adjust> 39 | \begin{code} 40 | _⊢_ : (A → B) → (B → Set) → (A → Set) 41 | (f ⊢ P) x = P (f x) 42 | \end{code} 43 | % 44 | \begin{code} 45 | 46 | 47 | \end{code} 48 | %<*bottom> 49 | \begin{code} 50 | data ⊥ : Set where 51 | \end{code} 52 | % 53 | \begin{code} 54 | 55 | \end{code} 56 | %<*sigma> 57 | \begin{code} 58 | record Σ (A : Set) (B : A → Set) : Set where 59 | constructor _,_ 60 | field proj₁ : A 61 | proj₂ : B proj₁ 62 | \end{code} 63 | % 64 | \begin{code} 65 | 66 | \end{code} 67 | %<*exists> 68 | \begin{code} 69 | ∃⟨_⟩ : (A → Set) → Set 70 | ∃⟨_⟩ P = Σ _ P 71 | \end{code} 72 | % 73 | \begin{code} 74 | 75 | \end{code} 76 | %<*dec> 77 | \begin{code} 78 | data Dec (P : Set) : Set where 79 | yes : P → Dec P 80 | no : (P → ⊥) → Dec P 81 | \end{code} 82 | % 83 | \begin{code} 84 | 85 | private 86 | variable 87 | a : A 88 | as : List A 89 | 90 | \end{code} 91 | %<*all> 92 | \begin{code} 93 | data All (P : A → Set) : List A → Set where 94 | [] : All P [] 95 | _∷_ : P a → All P as → All P (a ∷ as) 96 | \end{code} 97 | % 98 | -------------------------------------------------------------------------------- /doc/type-scope-semantics.agda/Semantics/Syntactic/Instances.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | module Semantics.Syntactic.Instances where 3 | 4 | open import Data.List.Base using (List; []; _∷_) 5 | open import Data.Var hiding (_<$>_) 6 | open import Data.Environment 7 | open import Syntax.Type 8 | open import Syntax.Calculus 9 | open import Semantics.Specification 10 | open import Semantics.Syntactic.Specification 11 | open import Relation.Unary 12 | open import Function 13 | 14 | private 15 | variable 16 | Γ Δ : List Type 17 | σ τ : Type 18 | 19 | open Syntactic 20 | 21 | -- We are, once more, using copatterns to prevent too much unfolding. 22 | 23 | \end{code} 24 | %<*synren> 25 | \begin{code} 26 | Syn^Ren : Syntactic Var 27 | Syn^Ren .zro = z 28 | Syn^Ren .th^𝓣 = th^Var 29 | Syn^Ren .var = `var 30 | \end{code} 31 | % 32 | %<*ren> 33 | \begin{code} 34 | Renaming : Semantics Var Term 35 | Renaming = syntactic Syn^Ren 36 | 37 | th^Term : Thinnable (Term σ) 38 | th^Term t ρ = semantics Renaming ρ t 39 | \end{code} 40 | % 41 | \begin{code} 42 | ren : Thinning Γ Δ → Term σ Γ → Term σ Δ 43 | ren ρ t = th^Term t ρ 44 | \end{code} 45 | %<*synsub> 46 | \begin{code} 47 | Syn^Sub : Syntactic Term 48 | Syn^Sub .zro = `var z 49 | Syn^Sub .th^𝓣 = th^Term 50 | Syn^Sub .var = id 51 | \end{code} 52 | % 53 | %<*sub> 54 | \begin{code} 55 | Substitution : Semantics Term Term 56 | Substitution = syntactic Syn^Sub 57 | 58 | sub : (Γ ─Env) Term Δ → Term σ Γ → Term σ Δ 59 | sub ρ t = semantics Substitution ρ t 60 | \end{code} 61 | % 62 | %<*eta> 63 | \begin{code} 64 | eta : ∀[ Term (σ `→ τ) ⇒ Term (σ `→ τ) ] 65 | eta t = `lam (`app (th^Term t extend) (`var z)) 66 | \end{code} 67 | % 68 | %<*betared> 69 | \begin{code} 70 | _⟨_/0⟩ : ∀[ (σ ∷_) ⊢ Term τ ⇒ Term σ ⇒ Term τ ] 71 | t ⟨ u /0⟩ = sub ((`var <$> identity) ∙ u) t 72 | \end{code} 73 | % 74 | -------------------------------------------------------------------------------- /doc/generic-syntax.agda/Generic/Simulation/Syntactic.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | {-# OPTIONS --safe --sized-types #-} 3 | 4 | module Generic.Simulation.Syntactic where 5 | 6 | open import Size 7 | open import Data.List hiding (lookup) 8 | open import Function 9 | open import Relation.Binary.PropositionalEquality 10 | 11 | open import Data.Var.Varlike 12 | open import Data.Environment 13 | open import Data.Relation as Relation 14 | open import Generic.Syntax 15 | open import Generic.Semantics 16 | open import Generic.Semantics.Syntactic 17 | open import Generic.Relator as Relator 18 | open import Generic.Simulation as Simulation 19 | open Simulation.Simulation 20 | 21 | module _ {I : Set} {d : Desc I} where 22 | 23 | \end{code} 24 | %<*renext> 25 | \begin{code} 26 | RenExt : Simulation d Ren Ren Eqᴿ Eqᴿ 27 | RenExt .thᴿ = λ ρ → cong (lookup ρ) 28 | RenExt .varᴿ = cong `var 29 | RenExt .algᴿ = λ _ _ → 30 | cong `con ∘ Relator.reifyᴿ Eqᴿ d (Simulation.reifyᴿ Eqᴿ Eqᴿ (vl^Refl vl^Var)) 31 | \end{code} 32 | % 33 | %<*subext> 34 | \begin{code} 35 | SubExt : Simulation d Sub Sub Eqᴿ Eqᴿ 36 | SubExt .thᴿ = λ ρ → cong (ren ρ) 37 | SubExt .varᴿ = id 38 | SubExt .algᴿ = λ _ _ → 39 | cong `con ∘ Relator.reifyᴿ Eqᴿ d (Simulation.reifyᴿ Eqᴿ Eqᴿ (vl^Refl vl^Tm)) 40 | \end{code} 41 | % 42 | %<*rensub> 43 | \begin{code} 44 | RenSub : Simulation d Ren Sub VarTmᴿ Eqᴿ 45 | \end{code} 46 | % 47 | \begin{code} 48 | RenSub .varᴿ = id 49 | RenSub .thᴿ = λ ρ → cong (λ t → th^Tm t ρ) 50 | RenSub .algᴿ = λ _ _ → 51 | cong `con ∘ Relator.reifyᴿ VarTmᴿ d (Simulation.reifyᴿ VarTmᴿ Eqᴿ vl^VarTm) 52 | 53 | private 54 | variable 55 | Γ Δ : List I 56 | σ : I 57 | 58 | \end{code} 59 | %<*rensubfun> 60 | \begin{code} 61 | rensub : (ρ : Thinning Γ Δ) (t : Tm d ∞ σ Γ) → ren ρ t ≡ sub (`var <$> ρ) t 62 | rensub ρ = Simulation.sim RenSub (packᴿ λ _ → refl) 63 | \end{code} 64 | % 65 | -------------------------------------------------------------------------------- /doc/type-scope-semantics.agda/Syntax/WeakHead.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | module Syntax.WeakHead where 3 | 4 | open import Syntax.Type 5 | open import Syntax.Calculus 6 | open import Semantics.Syntactic.Instances 7 | 8 | open import Data.Var 9 | open import Data.Environment 10 | open import Data.List.Base using (List; []; _∷_) 11 | 12 | open import Relation.Unary 13 | 14 | private 15 | 16 | variable 17 | 18 | σ τ : Type 19 | Γ Δ : List Type 20 | 21 | \end{code} 22 | %<*weakhead> 23 | \begin{code} 24 | data WHNE : Type ─Scoped where 25 | `var : ∀[ Var σ ⇒ WHNE σ ] 26 | `app : ∀[ WHNE (σ `→ τ) ⇒ Term σ ⇒ WHNE τ ] 27 | `ifte : ∀[ WHNE `Bool ⇒ Term σ ⇒ Term σ ⇒ WHNE σ ] 28 | 29 | data WHNF : Type ─Scoped where 30 | `lam : ∀[ (σ ∷_) ⊢ Term τ ⇒ WHNF (σ `→ τ) ] 31 | `one : ∀[ WHNF `Unit ] 32 | `tt `ff : ∀[ WHNF `Bool ] 33 | `whne : ∀[ WHNE σ ⇒ WHNF σ ] 34 | \end{code} 35 | % 36 | %<*thweakhead> 37 | \begin{code} 38 | th^WHNE : Thinnable (WHNE σ) 39 | th^WHNE (`var v) ρ = `var (lookup ρ v) 40 | th^WHNE (`app f t) ρ = `app (th^WHNE f ρ) (th^Term t ρ) 41 | th^WHNE (`ifte b l r) ρ = `ifte (th^WHNE b ρ) (th^Term l ρ) (th^Term r ρ) 42 | 43 | th^WHNF : Thinnable (WHNF σ) 44 | th^WHNF (`lam b) ρ = `lam (th^Term b (th^Env th^Var ρ extend ∙ z)) 45 | th^WHNF `one ρ = `one 46 | th^WHNF `tt ρ = `tt 47 | th^WHNF `ff ρ = `ff 48 | th^WHNF (`whne t) ρ = `whne (th^WHNE t ρ) 49 | \end{code} 50 | % 51 | %<*erase> 52 | \begin{code} 53 | erase^WHNE : ∀[ WHNE σ ⇒ Term σ ] 54 | erase^WHNE (`var v) = `var v 55 | erase^WHNE (`app f t) = `app (erase^WHNE f) t 56 | erase^WHNE (`ifte b l r) = `ifte (erase^WHNE b) l r 57 | 58 | erase^WHNF : ∀[ WHNF σ ⇒ Term σ ] 59 | erase^WHNF (`lam b) = `lam b 60 | erase^WHNF `one = `one 61 | erase^WHNF `tt = `tt 62 | erase^WHNF `ff = `ff 63 | erase^WHNF (`whne t) = erase^WHNE t 64 | \end{code} 65 | % 66 | -------------------------------------------------------------------------------- /doc/type-scope-semantics.agda/Syntax/Normal/Thinnable.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | module Syntax.Normal.Thinnable where 3 | 4 | open import Data.List.Base using (List; []; _∷_) 5 | open import Syntax.Type 6 | open import Syntax.Normal 7 | open import Data.Var 8 | open import Data.Environment 9 | open import Semantics.Syntactic.Instances 10 | open import Function 11 | open import Relation.Binary.PropositionalEquality 12 | 13 | private 14 | 15 | variable 16 | 17 | R : Type → Set 18 | σ : Type 19 | 20 | {- 21 | mutual 22 | 23 | wk^nf-trans′ : 24 | ∀ {Γ Δ Θ σ R inc₁ inc₃} {inc₂ : Renaming Δ Θ} → 25 | `∀[ Equality ] (Env.trans inc₁ inc₂) inc₃ → 26 | (t : Γ ⊢[ R ]^nf σ) → wk^nf inc₂ (wk^nf inc₁ t) ≡ wk^nf inc₃ t 27 | wk^nf-trans′ eq (`neu pr t) = cong (`neu pr) $ wk^ne-trans′ eq t 28 | wk^nf-trans′ eq `⟨⟩ = refl 29 | wk^nf-trans′ eq `tt = refl 30 | wk^nf-trans′ eq `ff = refl 31 | wk^nf-trans′ eq (`λ t) = 32 | cong `λ $ wk^nf-trans′ ((cong 1+_ ∘ lookup^R eq) ∙^R′ refl) t 33 | 34 | wk^ne-trans′ : 35 | ∀ {Γ Δ Θ σ R inc₁ inc₃} {inc₂ : Renaming Δ Θ} → 36 | `∀[ Equality ] (Env.trans inc₁ inc₂) inc₃ → 37 | (t : Γ ⊢[ R ]^ne σ) → wk^ne inc₂ (wk^ne inc₁ t) ≡ wk^ne inc₃ t 38 | wk^ne-trans′ eq (`var v) = cong `var $ lookup^R eq v 39 | wk^ne-trans′ eq (t `$ u) = cong₂ _`$_ (wk^ne-trans′ eq t) (wk^nf-trans′ eq u) 40 | wk^ne-trans′ eq (`ifte t l r) = 41 | cong₂ (uncurry `ifte) (cong₂ _,_ (wk^ne-trans′ eq t) (wk^nf-trans′ eq l)) 42 | (wk^nf-trans′ eq r) 43 | 44 | wk^nf-trans : 45 | ∀ {Γ Δ Θ σ R inc₁} {inc₂ : Renaming Δ Θ} → 46 | (t : Γ ⊢[ R ]^nf σ) → wk^nf inc₂ (wk^nf inc₁ t) ≡ wk^nf (Env.trans inc₁ inc₂) t 47 | wk^nf-trans = wk^nf-trans′ $ pack^R $ λ _ → refl 48 | 49 | wk^ne-trans : 50 | ∀ {Γ Δ Θ σ R inc₁} {inc₂ : Renaming Δ Θ} → 51 | (t : Γ ⊢[ R ]^ne σ) → wk^ne inc₂ (wk^ne inc₁ t) ≡ wk^ne (Env.trans inc₁ inc₂) t 52 | wk^ne-trans = wk^ne-trans′ $ pack^R $ λ _ → refl 53 | -} 54 | \end{code} 55 | -------------------------------------------------------------------------------- /doc/acknowledgments.tex: -------------------------------------------------------------------------------- 1 | \section*{Acknowledgements} 2 | 3 | I would like to thank my supervisor Conor McBride for his contagious 4 | enthusiasm, for his generosity, and most of all for the near limitless 5 | freedom I was given to pursue the problems I was interested in. 6 | % 7 | I learned a lot of the techniques used in this thesis through our 8 | long discussions. 9 | % 10 | Many thanks naturally go to my other co-authors. 11 | % 12 | James Chapman, James McKinna, and Bob Atkey helped me simplify my 13 | generic setup and clarify some of the categorical structures 14 | underlying this work. 15 | % 16 | Andreas Abel, Steven Schäfer, Kathrin Stark, Aliya Hameer, 17 | Alberto Momigliano, and Brigitte Pientka offered, 18 | via the POPLMark Reloaded Challenge, a great opportunity to 19 | compare, contrast, and mutually enrich our respective approaches 20 | to well scoped syntaxes with binding. 21 | 22 | I would like to thank my examiners Neel Krishnaswami and Jules Hedges 23 | for the detailed feedback that helped make this thesis a better document, 24 | and for a very pleasant and interesting discussion during the viva. 25 | 26 | The results in this thesis could not have been obtained without the 27 | tireless work of the research team behind Agda (Ulf Norell, 28 | Andreas Abel, Nils Anders Danielsson, Jesper Cockx). 29 | % 30 | I have learned so much from using the system, and have had a fantastic 31 | time contributing back to it and collaborating on new features during 32 | the regular Agda Implementors' Meetings. 33 | 34 | My work benefited immensely from Matthew Daggit's development and 35 | maintenance work for the Agda standard library. 36 | Our constant discussions on library issues 37 | have deeply informed my own design sensibilities. 38 | 39 | This thesis was a long time in the making. 40 | % 41 | I am immensely grateful to both Bart Jacobs and Edwin Brady for trusting 42 | me to conduct exciting postdoctoral research under their supervision 43 | while I had not yet defended. 44 | % 45 | This also means that I have made many wonderful friends in my time in 46 | Strathclyde, Nijmegen, and St Andrews and I want to conclude 47 | by thanking them for making these years very pleasant. 48 | 49 | Last but not least, I want to thank my parents who always supported my 50 | choices. 51 | -------------------------------------------------------------------------------- /doc/generic-syntax.agda/Generic/Examples/Printing.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | {-# OPTIONS --safe --sized-types #-} 3 | 4 | module Generic.Examples.Printing where 5 | 6 | open import Data.Bool 7 | open import Data.Product 8 | open import Data.List.Base using (_∷_; []) 9 | open import Data.String 10 | open import Function 11 | open import Relation.Binary.PropositionalEquality 12 | 13 | open import Data.Var 14 | open import Data.Environment 15 | open import Generic.Syntax 16 | open import Semantics.Printing using (getW) 17 | open import Generic.Semantics.Printing 18 | 19 | -- UTLC 20 | open import Generic.Syntax.UTLC 21 | 22 | \end{code} 23 | %<*printUTLC> 24 | \begin{code} 25 | printUTLC : Display UTLC 26 | printUTLC = λ where 27 | (`app' f t) → f ++ " (" ++ t ++ ")" 28 | (`lam' (x , b)) → "λ" ++ getW (lookup x z) ++ ". " ++ b 29 | \end{code} 30 | % 31 | \begin{code} 32 | 33 | open import Agda.Builtin.Equality 34 | 35 | \end{code} 36 | %<*printid> 37 | \begin{code} 38 | _ : print printUTLC id^U ≡ "λa. a" 39 | _ = refl 40 | \end{code} 41 | % 42 | %<*printopen> 43 | \begin{code} 44 | _ : let tm : Tm UTLC _ _ (_ ∷ _ ∷ []) 45 | tm = `app (`var z) (`lam (`var (s (s z)))) 46 | in print printUTLC tm ≡ "b (λc. a)" 47 | _ = refl 48 | \end{code} 49 | % 50 | \begin{code} 51 | 52 | open import Generic.Examples.HuttonsRazor 53 | 54 | open import Data.Nat.Show as Nat 55 | 56 | printHR : Display HuttRaz 57 | printHR = λ where 58 | (true , n , _) → Nat.show n 59 | (false , m , n , _) → m ++ " + " ++ n 60 | 61 | _ : let expr : TM HuttRaz _ ;expr = ((lit 2 [+] lit 6) [+] lit 0) 62 | in print printHR expr ≡ "2 + 6 + 0" 63 | _ = refl 64 | 65 | open import Generic.Examples.SystemF as SystemF 66 | 67 | printSF : Display system-F 68 | printSF = case (λ { (σ , τ , _) → σ ++ " → " ++ τ }) 69 | $ case (λ { ((x , b) , _) → "∀" ++ getW (lookup x z) ++ ". " ++ b }) 70 | $ case (λ { (f , t , _) → f ++ "(" ++ t ++ ")" }) 71 | $ case (λ { ((x , b) , _) → "λ" ++ getW (lookup x z) ++ ". " ++ b }) 72 | $ case (λ { (f , T , _) → f ++ "(" ++ T ++ ")" }) 73 | (λ { ((x , b) , _) → "Λ" ++ getW (lookup x z) ++ ". " ++ b }) 74 | 75 | _ : print printSF SystemF.`id ≡ "Λa. λb. b" 76 | _ = refl 77 | \end{code} 78 | -------------------------------------------------------------------------------- /doc/generic-syntax.tex/other-programs.tex: -------------------------------------------------------------------------------- 1 | \chapter{Other Programs} 2 | 3 | All the programs operating on syntax we have seen so far have been 4 | instances of our notion of \AR{Semantics}. This shows that our 5 | framework covers a lot of useful cases. 6 | 7 | An interesting question raised by this observation is whether there 8 | are any interesting traversals that are not captured by this setting. 9 | At least two types of programs are clearly not instances of \AR{Semantics}: 10 | those that do not assign a meaning to a term in terms of the meaning of 11 | its direct subterms, and those whose return type depends on the input 12 | term. 13 | 14 | In this chapter we study two such programs. We also demonstrate that 15 | although they may not fit the exact pattern we have managed to abstract, 16 | it is still sometimes possible to take advantage of our data-generic 17 | setting and to implement them once and for all syntaxes with binding. 18 | 19 | \section{Big Step Evaluators} 20 | 21 | Chapman's thesis~(\citeyear{chapman2009type}) provides us with a good 22 | example of a function that does not assign a meaning to a term by only 23 | relying on the meaning of its direct subterms: a big step evaluator. 24 | 25 | Setting aside the issue of proving such a function terminating, a big 26 | step evaluator computes the normal form of a term by first recursively 27 | computing the normal forms of its subterms and then either succeeding 28 | or, having uncovered a new redex, firing it and then normalising the 29 | reduct. This is embodied by the case for application 30 | (cf. ~\cref{figure:bigstepapp}). 31 | 32 | \begin{figure} 33 | \begin{mathpar} 34 | \inferrule 35 | {f ⇓ λx.b \and t ⇓ v \and b [x ↦ v] ⇓ nf} 36 | {f\,t ⇓ nf} 37 | \end{mathpar} 38 | \caption{Big step evaluation of an application}\label{figure:bigstepapp} 39 | \end{figure} 40 | 41 | In contrast, the \AR{Semantics} constraints enforce that the meaning 42 | of each term constructor is directly expressed in terms of the meaning 43 | of its arguments. There is no way to ``restart'' the evaluator once the 44 | redex has been fired. This means we will not be able to make it an 45 | instance of our framework and, consequently, we will not be able to use 46 | the tools introduced in the next chapters to reason generically over 47 | such functions. 48 | 49 | \input{generic-syntax.tex/catalogue/equality.tex} 50 | -------------------------------------------------------------------------------- /doc/generic-syntax.agda/Generic/Semantics/NbyE.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | {-# OPTIONS --sized-types #-} 3 | module Generic.Semantics.NbyE where 4 | 5 | open import Size 6 | open import Data.Unit 7 | open import Data.Bool 8 | open import Data.Product 9 | open import Data.List.Base hiding ([_]) 10 | open import Function 11 | open import Relation.Unary 12 | open import Relation.Binary.PropositionalEquality hiding ([_]) 13 | 14 | open import Data.Var using (Var; _─Scoped) 15 | open import Data.Var.Varlike 16 | open import Data.Environment hiding (_<$>_; sequenceA) 17 | open import Generic.Syntax 18 | open import Generic.Semantics 19 | 20 | private 21 | variable 22 | I : Set 23 | s : Size 24 | σ : I 25 | d : Desc I 26 | 27 | \end{code} 28 | %<*domain> 29 | \begin{code} 30 | {-# NO_POSITIVITY_CHECK #-} 31 | data Dm (d : Desc I) : Size → I ─Scoped where 32 | V : ∀[ Var σ ⇒ Dm d s σ ] 33 | C : ∀[ ⟦ d ⟧ (Kripke (Dm d s) (Dm d s)) σ ⇒ Dm d (↑ s) σ ] 34 | ⊥ : ∀[ Dm d (↑ s) σ ] 35 | \end{code} 36 | % 37 | \begin{code} 38 | module _ {d : Desc I} where 39 | 40 | th^Dm : Thinnable (Dm d s σ) 41 | th^Dm (V k) ρ = V (th^Var k ρ) 42 | th^Dm (C t) ρ = C (fmap d (λ Θ i kr → th^Kr Θ th^Dm kr ρ) t) 43 | th^Dm ⊥ ρ = ⊥ 44 | 45 | vl^Dm : VarLike (Dm d s) 46 | vl^Dm = record { new = V Var.z ; th^𝓥 = th^Dm } 47 | 48 | 49 | open import Data.Maybe.Base 50 | import Data.Maybe.Categorical as Maybe 51 | import Category.Monad as CM 52 | open import Level 53 | private module M = CM.RawMonad (Maybe.monad {zero}) 54 | instance _ = M.rawIApplicative 55 | open M 56 | 57 | Alg : (d : Desc I) (𝓥 𝓒 : I ─Scoped) → Set 58 | Alg d 𝓥 𝓒 = ∀ {σ Γ} → ⟦ d ⟧ (Kripke 𝓥 𝓒) σ Γ → 𝓒 σ Γ 59 | 60 | module _ {I : Set} {d : Desc I} where 61 | \end{code} 62 | %<*nbe-setup> 63 | \begin{code} 64 | reify^Dm : ∀[ Dm d s σ ⇒ Maybe ∘ Tm d ∞ σ ] 65 | nbe : Alg d (Dm d ∞) (Dm d ∞) → Semantics d (Dm d ∞) (Dm d ∞) 66 | 67 | norm : Alg d (Dm d ∞) (Dm d ∞) → ∀[ Tm d ∞ σ ⇒ Maybe ∘ Tm d ∞ σ ] 68 | norm alg = reify^Dm ∘ Semantics.semantics (nbe alg) (base vl^Dm) 69 | \end{code} 70 | % 71 | \begin{code} 72 | reify^Dm (V k) = just (`var k) 73 | reify^Dm (C v) = `con <$> mapA d (λ Θ i → reify^Dm ∘ reify vl^Dm Θ i) v 74 | reify^Dm ⊥ = nothing 75 | 76 | Semantics.th^𝓥 (nbe alg) = th^Dm 77 | Semantics.var (nbe alg) = id 78 | Semantics.alg (nbe alg) = alg 79 | \end{code} 80 | -------------------------------------------------------------------------------- /doc/type-scope-semantics.agda/StateOfTheArt/McBride05/Kit.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | module StateOfTheArt.McBride05.Kit where 3 | 4 | open import Data.List hiding ([_]; lookup) 5 | open import Data.Var hiding (_<$>_) 6 | open import Data.Environment 7 | open import Relation.Unary 8 | open import Function 9 | open import StateOfTheArt.McBride05.Base using (Type; Tm) 10 | open Type 11 | open Tm 12 | 13 | private 14 | variable 15 | σ τ : Type 16 | Γ Δ : List Type 17 | ⧫ : Type ─Scoped 18 | 19 | \end{code} 20 | %<*kitdef> 21 | \begin{code} 22 | record Kit (⧫ : Type ─Scoped) : Set where 23 | field var : ∀[ ⧫ σ ⇒ Tm σ ] 24 | zro : ∀[ (σ ∷_) ⊢ ⧫ σ ] 25 | wkn : ∀[ ⧫ τ ⇒ (σ ∷_) ⊢ ⧫ τ ] 26 | \end{code} 27 | % 28 | \begin{code} 29 | 30 | open Kit 31 | 32 | \end{code} 33 | %<*kitsem> 34 | \begin{code} 35 | kit : Kit ⧫ → (Γ ─Env) ⧫ Δ → Tm σ Γ → Tm σ Δ 36 | kit K ρ (`var v) = K .var (lookup ρ v) 37 | kit K ρ (`app f t) = `app (kit K ρ f) (kit K ρ t) 38 | kit K ρ (`lam b) = `lam (kit K ρ′ b) 39 | where ρ′ = (K .wkn <$> ρ) ∙ K .zro 40 | \end{code} 41 | % 42 | 43 | %<*renkit> 44 | \begin{code} 45 | ren^Kit : Kit Var 46 | ren^Kit .var = `var 47 | ren^Kit .zro = z 48 | ren^Kit .wkn = s 49 | \end{code} 50 | % 51 | 52 | %<*ren> 53 | \begin{code} 54 | ren : (Γ ─Env) Var Δ → Tm σ Γ → Tm σ Δ 55 | ren = kit ren^Kit 56 | \end{code} 57 | % 58 | 59 | %<*subkit> 60 | \begin{code} 61 | sub^Kit : Kit Tm 62 | sub^Kit .var = id 63 | sub^Kit .zro = `var z 64 | sub^Kit .wkn = ren (pack s) 65 | \end{code} 66 | % 67 | %<*sub> 68 | \begin{code} 69 | sub : (Γ ─Env) Tm Δ → Tm σ Γ → Tm σ Δ 70 | sub = kit sub^Kit 71 | \end{code} 72 | % 73 | \begin{code} 74 | 75 | 76 | Model : Type ─Scoped 77 | Model base Γ = Tm base Γ 78 | Model (arr σ τ) Γ = ∀ {Δ} → Thinning Γ Δ → Model σ Δ → Model τ Δ 79 | 80 | th^Model : ∀ {σ} → Thinning Γ Δ → Model σ Γ → Model σ Δ 81 | th^Model {σ = base } ρ v = ren ρ v 82 | th^Model {σ = arr σ τ} ρ v = v ∘ (select ρ) 83 | 84 | APP : Model (arr σ τ) Γ → Model σ Γ → Model τ Γ 85 | APP f t = f (pack id) t 86 | 87 | LAM : Model (arr σ τ) Γ → Model (arr σ τ) Γ 88 | LAM = id 89 | \end{code} 90 | %<*nbe> 91 | \begin{code} 92 | nbe : (Γ ─Env) Model Δ → Tm σ Γ → Model σ Δ 93 | nbe ρ (`var v) = lookup ρ v 94 | nbe ρ (`app f t) = APP (nbe ρ f) (nbe ρ t) 95 | nbe ρ (`lam t) = LAM (λ re v → nbe ((th^Model re <$> ρ) ∙ v) t) 96 | \end{code} 97 | % 98 | -------------------------------------------------------------------------------- /doc/shared/Data/Var.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | {-# OPTIONS --safe #-} 3 | 4 | -- When scopes are represented by lists of kinds, a variable of 5 | -- a given kind is a position in such a list. This is a strongly 6 | -- typed version of de Bruijn indices hence the name we picked 7 | -- for Var's constructor: 8 | -- * z for zero 9 | -- * s for successor 10 | 11 | module Data.Var where 12 | 13 | open import Data.Sum hiding (map) 14 | open import Data.List.Base hiding ([_]; _─_) 15 | open import Data.List.Relation.Unary.All using (All ; _∷_) 16 | open import Relation.Unary 17 | open import Function.Base 18 | open import Agda.Builtin.Equality 19 | 20 | \end{code} 21 | %<*scoped> 22 | \begin{code} 23 | _─Scoped : Set → Set₁ 24 | I ─Scoped = I → List I → Set 25 | \end{code} 26 | % 27 | \begin{code} 28 | 29 | private 30 | variable 31 | I J : Set 32 | σ τ : I 33 | 34 | \end{code} 35 | %<*var> 36 | \begin{code} 37 | data Var : I ─Scoped where 38 | z : ∀[ (σ ∷_) ⊢ Var σ ] 39 | s : ∀[ Var σ ⇒ (τ ∷_) ⊢ Var σ ] 40 | \end{code} 41 | % 42 | \begin{code} 43 | infixl 3 _─_ 44 | _─_ : {i : I} (Γ : List I) → Var i Γ → List I 45 | _ ∷ Γ ─ z = Γ 46 | σ ∷ Γ ─ s v = σ ∷ (Γ ─ v) 47 | 48 | get : {B : I → Set} → ∀[ Var σ ⇒ All B ⇒ const (B σ) ] 49 | get z (b ∷ _) = b 50 | get (s v) (_ ∷ bs) = get v bs 51 | 52 | _<$>_ : (f : I → J) → ∀[ Var σ ⇒ map f ⊢ Var (f σ) ] 53 | f <$> z = z 54 | f <$> s v = s (f <$> v) 55 | 56 | record Injective (f : I → J) : Set where 57 | constructor mkInjective 58 | field inj : ∀ {i₁ i₂} → f i₁ ≡ f i₂ → i₁ ≡ i₂ 59 | open Injective public 60 | 61 | Injective-inj₁ : ∀ {A B : Set} → Injective ((A → A ⊎ B) ∋ inj₁) 62 | inj Injective-inj₁ refl = refl 63 | 64 | Injective-inj₂ : ∀ {A B : Set} → Injective ((B → A ⊎ B) ∋ inj₂) 65 | inj Injective-inj₂ refl = refl 66 | 67 | _<$>⁻¹_ : {f : I → J} → Injective f → 68 | ∀[ map f ⊢ Var (f σ) ⇒ Var σ ] 69 | _<$>⁻¹_ {f = f} F = go _ refl refl where 70 | 71 | go : ∀ {j} Γ {js} → f σ ≡ j → map f Γ ≡ js → Var j js → Var σ Γ 72 | go [] eq () z 73 | go [] eq () (s v) 74 | go (i ∷ is) eq refl z rewrite inj F eq = z 75 | go (i ∷ is) eq refl (s v) = s (go is eq refl v) 76 | 77 | injectˡ : ∀ Δ → ∀[ Var σ ⇒ (_++ Δ) ⊢ Var σ ] 78 | injectˡ k z = z 79 | injectˡ k (s v) = s (injectˡ k v) 80 | 81 | injectʳ : ∀ Δ → ∀[ Var σ ⇒ (Δ ++_) ⊢ Var σ ] 82 | injectʳ [] v = v 83 | injectʳ (y ∷ ys) v = s (injectʳ ys v) 84 | \end{code} 85 | -------------------------------------------------------------------------------- /doc/type-scope-semantics.agda/Semantics/CPS/CBV.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | module Semantics.CPS.CBV where 3 | 4 | open import Syntax.Type 5 | open import Syntax.Calculus 6 | open import Syntax.MoggiML.Type 7 | open import Syntax.MoggiML.Calculus 8 | 9 | open import Data.List.Base using (List; []; _∷_; map) 10 | open import Data.Product as Prod hiding (map) 11 | open import Data.Var 12 | open import Data.Environment hiding (_<$>_; uncurry) 13 | open import Semantics.Specification 14 | open import Function 15 | open import Relation.Unary 16 | open import Relation.Binary.PropositionalEquality 17 | 18 | private 19 | 20 | variable 21 | 22 | σ τ : Type 23 | \end{code} 24 | %<*cbv> 25 | \begin{code} 26 | mutual 27 | 28 | #CBV : Type → CType 29 | #CBV σ = # (CBV σ) 30 | 31 | CBV : Type → CType 32 | CBV `Unit = `Unit 33 | CBV `Bool = `Bool 34 | CBV (σ `→ τ) = CBV σ `→# CBV τ 35 | \end{code} 36 | % 37 | \begin{code} 38 | CBV-inj : ∀ σ τ → CBV σ ≡ CBV τ → σ ≡ τ 39 | CBV-inj `Unit `Unit _ = refl 40 | CBV-inj `Unit `Bool () 41 | CBV-inj `Unit (_ `→ _) () 42 | CBV-inj `Bool `Unit () 43 | CBV-inj `Bool `Bool _ = refl 44 | CBV-inj `Bool (_ `→ _) () 45 | CBV-inj (_ `→ _) `Unit () 46 | CBV-inj (_ `→ _) `Bool () 47 | CBV-inj (σ₁ `→ τ₁) (σ₂ `→ τ₂) eq = 48 | uncurry (cong₂ _`→_) (Prod.map (CBV-inj σ₁ σ₂) (CBV-inj τ₁ τ₂) (`→#-inj eq)) 49 | \end{code} 50 | %<*cbvalue> 51 | \begin{code} 52 | V^CBV : Type ─Scoped 53 | V^CBV σ Γ = Var (CBV σ) (map CBV Γ) 54 | 55 | C^CBV : Type ─Scoped 56 | C^CBV σ Γ = ML (# CBV σ) (map CBV Γ) 57 | \end{code} 58 | % 59 | \begin{code} 60 | 61 | open Semantics 62 | \end{code} 63 | %<*thinnable> 64 | \begin{code} 65 | th^V : Thinnable (V^CBV σ) 66 | th^V v ρ = CBV <$> th^Var (mkInjective (CBV-inj _ _) <$>⁻¹ v) ρ 67 | \end{code} 68 | % 69 | %<*lam> 70 | \begin{code} 71 | LAM : ∀[ □ (V^CBV σ ⇒ C^CBV τ) ⇒ C^CBV (σ `→ τ) ] 72 | LAM b = `ret (`lam (b extend z)) 73 | \end{code} 74 | % 75 | %<*app> 76 | \begin{code} 77 | APP : ∀[ C^CBV (σ `→ τ) ⇒ C^CBV σ ⇒ C^CBV τ ] 78 | APP f t = `bind f (`lam (`bind (th^ML t extend) (`lam (`app (`var (s z)) (`var z))))) 79 | \end{code} 80 | % 81 | %<*ifte> 82 | \begin{code} 83 | IFTE : ∀[ C^CBV `Bool ⇒ C^CBV σ ⇒ C^CBV σ ⇒ C^CBV σ ] 84 | IFTE b l r = `bind b (`lam (`ifte (`var z) (th^ML l extend) (th^ML r extend))) 85 | \end{code} 86 | % 87 | %<*eval> 88 | \begin{code} 89 | CPS : Semantics V^CBV C^CBV 90 | CPS .th^𝓥 = th^V 91 | CPS .var = `ret ∘ `var 92 | CPS .lam = LAM 93 | CPS .app = APP 94 | CPS .one = `ret `one 95 | CPS .tt = `ret `tt 96 | CPS .ff = `ret `ff 97 | CPS .ifte = IFTE 98 | \end{code} 99 | % 100 | -------------------------------------------------------------------------------- /doc/generic-syntax.agda/Generic/Equality.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | {-# OPTIONS --with-K --safe --sized-types #-} 3 | 4 | module Generic.Equality where 5 | 6 | open import Size 7 | open import Data.Unit 8 | open import Data.Product 9 | open import Data.List.Base using (List; []; _∷_) 10 | open import Relation.Nullary 11 | open import Relation.Binary.PropositionalEquality 12 | open import Relation.Binary.PropositionalEquality.WithK 13 | 14 | open import Data.Var 15 | open import Generic.Syntax 16 | 17 | private 18 | variable 19 | I : Set 20 | σ : I 21 | Γ : List I 22 | d : Desc I 23 | i : Size 24 | \end{code} 25 | \begin{code} 26 | 27 | \end{code} 28 | %<*constraints> 29 | \begin{code} 30 | Constraints : Desc I → Set 31 | Constraints (`σ A d) = ((a b : A) → Dec (a ≡ b)) × (∀ a → Constraints (d a)) 32 | Constraints (`X _ _ d) = Constraints d 33 | Constraints (`∎ _) = ⊤ 34 | \end{code} 35 | % 36 | \begin{code} 37 | 38 | \end{code} 39 | %<*eqVarType> 40 | \begin{code} 41 | eq^Var : (v w : Var σ Γ) → Dec (v ≡ w) 42 | \end{code} 43 | % 44 | %<*eqVarNo> 45 | \begin{code} 46 | eq^Var z (s w) = no (λ ()) 47 | eq^Var (s v) z = no (λ ()) 48 | \end{code} 49 | % 50 | %<*eqVarYesZ> 51 | \begin{code} 52 | eq^Var z z = yes refl 53 | \end{code} 54 | % 55 | %<*eqVarYesS> 56 | \begin{code} 57 | eq^Var (s v) (s w) with eq^Var v w 58 | ... | yes p = yes (cong s p) 59 | ... | no ¬p = no λ where refl → ¬p refl 60 | \end{code} 61 | % 62 | \begin{code} 63 | 64 | module _ (eq^d : Constraints d) where 65 | 66 | \end{code} 67 | %<*eqTmType> 68 | \begin{code} 69 | eq^Tm : (t u : Tm d i σ Γ) → Dec (t ≡ u) 70 | eq^⟦⟧ : ∀ e → Constraints e → (b c : ⟦ e ⟧ (Scope (Tm d i)) σ Γ) → Dec (b ≡ c) 71 | \end{code} 72 | % 73 | \begin{code} 74 | 75 | 76 | eq^Tm (`var v) (`con c) = no (λ ()) 77 | eq^Tm (`con b) (`var w) = no (λ ()) 78 | eq^Tm (`var v) (`var w) with eq^Var v w 79 | ... | yes p = yes (cong `var p) 80 | ... | no ¬p = no (λ where refl → ¬p refl) 81 | eq^Tm (`con b) (`con c) with eq^⟦⟧ _ eq^d b c 82 | ... | yes p = yes (cong `con p) 83 | ... | no ¬p = no (λ where refl → ¬p refl) 84 | 85 | eq^⟦⟧ (`σ A e) (eq^A , eq^e) (a₁ , b) (a₂ , c) 86 | with eq^A a₁ a₂ 87 | ... | no ¬p = no (λ where refl → ¬p refl) 88 | ... | yes refl 89 | with eq^⟦⟧ (e a₁) (eq^e a₁) b c 90 | ... | yes q = yes (cong (_ ,_) q) 91 | ... | no ¬q = no (λ where refl → ¬q refl) 92 | eq^⟦⟧ (`X _ _ e) eq^e (t , b) (u , c) 93 | with eq^Tm t u | eq^⟦⟧ e eq^e b c 94 | ... | yes p | yes q = yes (cong₂ _,_ p q) 95 | ... | no ¬p | _ = no (λ where refl → ¬p refl) 96 | ... | _ | no ¬q = no (λ where refl → ¬q refl) 97 | eq^⟦⟧ (`∎ _) eq^e b c = yes (≡-irrelevant b c) 98 | \end{code} 99 | -------------------------------------------------------------------------------- /doc/type-scope-semantics.agda/Semantics/CPS/CBN.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | module Semantics.CPS.CBN where 3 | 4 | open import Syntax.Type 5 | open import Syntax.Calculus 6 | open import Syntax.MoggiML.Type 7 | open import Syntax.MoggiML.Calculus 8 | 9 | open import Data.List.Base using (List; []; _∷_; map) 10 | open import Data.Product as Prod hiding (map) 11 | open import Data.Var 12 | open import Data.Environment hiding (_<$>_; uncurry) 13 | open import Semantics.Specification 14 | open import Function 15 | open import Relation.Unary 16 | open import Relation.Binary.PropositionalEquality 17 | 18 | private 19 | 20 | variable 21 | 22 | σ τ : Type 23 | 24 | \end{code} 25 | %<*cbn> 26 | \begin{code} 27 | mutual 28 | 29 | #CBN : Type → CType 30 | #CBN σ = # (CBN σ) 31 | 32 | CBN : Type → CType 33 | CBN `Unit = `Unit 34 | CBN `Bool = `Bool 35 | CBN (σ `→ τ) = #CBN σ `→# CBN τ 36 | \end{code} 37 | % 38 | \begin{code} 39 | CBN-inj : ∀ σ τ → CBN σ ≡ CBN τ → σ ≡ τ 40 | CBN-inj `Unit `Unit eq = refl 41 | CBN-inj `Unit `Bool () 42 | CBN-inj `Unit (_ `→ _) () 43 | CBN-inj `Bool `Unit () 44 | CBN-inj `Bool `Bool eq = refl 45 | CBN-inj `Bool (_ `→ _) () 46 | CBN-inj (_ `→ _) `Unit () 47 | CBN-inj (_ `→ _) `Bool () 48 | CBN-inj (σ₁ `→ τ₁) (σ₂ `→ τ₂) eq = 49 | uncurry (cong₂ _`→_) (Prod.map (CBN-inj σ₁ σ₂ ∘ #-inj) (CBN-inj τ₁ τ₂) (`→#-inj eq)) 50 | 51 | #CBN-inj : ∀ {σ τ} → #CBN σ ≡ #CBN τ → σ ≡ τ 52 | #CBN-inj = CBN-inj _ _ ∘ #-inj 53 | \end{code} 54 | %<*cbntransformer> 55 | \begin{code} 56 | _^CBN : CType ─Scoped → Type ─Scoped 57 | (T ^CBN) σ Γ = T (#CBN σ) (map #CBN Γ) 58 | \end{code} 59 | % 60 | %<*thinnable> 61 | \begin{code} 62 | th^V : Thinnable ((Var ^CBN) σ) 63 | th^V v ρ = #CBN <$> th^Var ((mkInjective #CBN-inj) <$>⁻¹ v) ρ 64 | \end{code} 65 | % 66 | %<*lam> 67 | \begin{code} 68 | LAM : ∀[ □ ((Var ^CBN) σ ⇒ (ML ^CBN) τ) ⇒ (ML ^CBN) (σ `→ τ) ] 69 | LAM b = `ret (`lam (b extend z)) 70 | \end{code} 71 | % 72 | %<*app> 73 | \begin{code} 74 | APP : ∀[ (ML ^CBN) (σ `→ τ) ⇒ (ML ^CBN) σ ⇒ (ML ^CBN) τ ] 75 | APP f t = `bind f (`lam (`app (`var z) (th^ML t extend))) 76 | \end{code} 77 | % 78 | %<*ifte> 79 | \begin{code} 80 | IFTE : ∀[ (ML ^CBN) `Bool ⇒ (ML ^CBN) σ ⇒ (ML ^CBN) σ ⇒ (ML ^CBN) σ ] 81 | IFTE b l r = `bind b (`lam (`ifte (`var z) (th^ML l extend) (th^ML r extend))) 82 | \end{code} 83 | % 84 | \begin{code} 85 | open Semantics 86 | 87 | \end{code} 88 | %<*eval> 89 | \begin{code} 90 | CPS : Semantics (Var ^CBN) (ML ^CBN) 91 | CPS .th^𝓥 = th^V 92 | CPS .var = `var 93 | CPS .lam = LAM 94 | CPS .app = APP 95 | CPS .one = `ret `one 96 | CPS .tt = `ret `tt 97 | CPS .ff = `ret `ff 98 | CPS .ifte = IFTE 99 | \end{code} 100 | % 101 | -------------------------------------------------------------------------------- /doc/type-scope-semantics.agda/Syntax/Calculus.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | {-# OPTIONS --safe --sized-types #-} 3 | 4 | module Syntax.Calculus where 5 | 6 | open import Data.Var 7 | open import Syntax.Type 8 | open import Data.List.Base using (List; []; _∷_) 9 | open import Relation.Unary 10 | 11 | private 12 | variable 13 | σ τ : Type 14 | 15 | -- The calculus is defined in a well-scoped and well-typed 16 | -- manner using an inductive family. A term effectively 17 | -- correponds to a derivation in the sequent calculus. 18 | 19 | -- Nota Bene: there are TWO proofs of Γ ⊢ `Bool corresponding 20 | -- to true and false respectively. 21 | 22 | module DISPLAYONLY where 23 | 24 | \end{code} 25 | %<*termcompact> 26 | %<*termcompact:decl> 27 | \begin{code} 28 | data Term : Type ─Scoped where 29 | \end{code} 30 | % 31 | \begin{code} 32 | `var : ∀[ Var σ ⇒ Term σ ] 33 | \end{code} 34 | %<*termcompact:lam> 35 | \begin{code} 36 | `lam : ∀[ (σ ∷_) ⊢ Term τ 37 | ⇒ Term (σ `→ τ) ] 38 | \end{code} 39 | % 40 | %<*termcompact:base> 41 | \begin{code} 42 | `one : ∀[ Term `Unit ] 43 | 44 | `tt : ∀[ Term `Bool ] 45 | `ff : ∀[ Term `Bool ] 46 | \end{code} 47 | % 48 | %<*termcompact:struct> 49 | \begin{code} 50 | `app : ∀[ Term (σ `→ τ) 51 | ⇒ Term σ 52 | ⇒ Term τ ] 53 | \end{code} 54 | %<*termcompact:ifte> 55 | \begin{code} 56 | `ifte : ∀[ Term `Bool 57 | ⇒ Term σ ⇒ Term σ 58 | ⇒ Term σ ] 59 | \end{code} 60 | % 61 | % 62 | 63 | 64 | \end{code} 65 | %<*term> 66 | \begin{code} 67 | data Term : Type ─Scoped where 68 | 69 | `var : ∀[ Var σ 70 | --------- 71 | ⇒ Term σ ] 72 | 73 | `app : ∀[ Term (σ `→ τ) ⇒ Term σ 74 | ---------------------- 75 | ⇒ Term τ ] 76 | 77 | `lam : ∀[ (σ ∷_) ⊢ Term τ 78 | --------------- 79 | ⇒ Term (σ `→ τ) ] 80 | 81 | `one : ∀[ 82 | --------- 83 | Term `Unit ] 84 | 85 | `tt : ∀[ 86 | --------- 87 | Term `Bool ] 88 | 89 | `ff : ∀[ 90 | --------- 91 | Term `Bool ] 92 | 93 | `ifte : ∀[ Term `Bool ⇒ Term σ ⇒ Term σ 94 | ---------------------------- 95 | ⇒ Term σ ] 96 | \end{code} 97 | % 98 | 99 | \begin{code} 100 | private 101 | \end{code} 102 | %<*id> 103 | \begin{code} 104 | id : ∀[ Term (σ `→ σ) ] 105 | id = `lam (`var z) 106 | \end{code} 107 | % 108 | 109 | %<*not> 110 | \begin{code} 111 | not : ∀[ Term (`Bool `→ `Bool) ] 112 | not = `lam (`ifte (`var z) `ff `tt) 113 | \end{code} 114 | % 115 | -------------------------------------------------------------------------------- /doc/generic-syntax.agda/Generic/Syntax/Bidirectional.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | {-# OPTIONS --safe --sized-types #-} 3 | 4 | module Generic.Syntax.Bidirectional where 5 | 6 | open import Size 7 | open import Data.Product 8 | open import Agda.Builtin.List 9 | open import Agda.Builtin.Equality 10 | open import Function 11 | 12 | open import Relation.Unary 13 | open import Data.Var 14 | open import Generic.Syntax 15 | 16 | infixr 5 _`→_ 17 | data Type : Set where 18 | α : Type 19 | _`→_ : Type → Type → Type 20 | 21 | -- We have an *untyped* language presented in a bidirectional manner 22 | -- where phases are statically checked 23 | 24 | \end{code} 25 | %<*mode> 26 | \begin{code} 27 | data Mode : Set where 28 | Check Synth : Mode 29 | \end{code} 30 | % 31 | %<*tag> 32 | \begin{code} 33 | data `Bidi : Set where 34 | App Lam Emb : `Bidi 35 | Cut : Type → `Bidi 36 | \end{code} 37 | % 38 | \begin{code} 39 | 40 | private 41 | variable 42 | m : Mode 43 | 44 | -- typesetting only 45 | module TYPESETTING-BIDI where 46 | \end{code} 47 | %<*databidi> 48 | \begin{code} 49 | data Bidi : Mode ─Scoped where 50 | `var : ∀[ Var m ⇒ Bidi m ] 51 | -- choice of four constructors: 52 | `app : ∀[ Bidi Synth ⇒ Bidi Check ⇒ 53 | Bidi Synth ] 54 | `lam : ∀[ (Synth ∷_) ⊢ Bidi Check ⇒ 55 | Bidi Check ] 56 | `cut : Type → ∀[ Bidi Check ⇒ Bidi Synth ] 57 | `emb : ∀[ Bidi Synth ⇒ Bidi Check ] 58 | \end{code} 59 | % 60 | 61 | -- On top of the traditional Application and Lambda-Abstraction constructors, 62 | -- we have two change of direction ones: `Emb` which takes inferable terms and 63 | -- makes them checkable (it is enough to compare the inferred type to the 64 | -- candidate provided); and `Cut` which takes a checkable term and makes it 65 | -- inferrable thanks to a type-annotation. 66 | 67 | \end{code} 68 | %<*desc> 69 | \begin{code} 70 | Bidi : Desc Mode 71 | Bidi = `σ `Bidi $ λ where 72 | -- var will be freely adjoined 73 | App → `X [] Synth (`X [] Check 74 | (`∎ Synth)) 75 | Lam → `X (Synth ∷ []) Check 76 | (`∎ Check) 77 | (Cut σ) → `X [] Check (`∎ Synth) 78 | Emb → `X [] Synth (`∎ Check) 79 | \end{code} 80 | % 81 | \begin{code} 82 | 83 | module PATTERNS where 84 | 85 | 86 | pattern `app' f t = (App , f , t , refl) 87 | pattern `lam' b = (Lam , b , refl) 88 | pattern `cut' σ t = (Cut σ , t , refl) 89 | pattern `emb' t = (Emb , t , refl) 90 | 91 | pattern `app f t = `con (`app' f t) 92 | pattern `lam b = `con (`lam' b) 93 | pattern `cut σ t = `con (`cut' σ t) 94 | pattern `emb t = `con (`emb' t) 95 | 96 | \end{code} 97 | 98 | %<*BDid> 99 | \begin{code} 100 | id^B : TM Bidi Check 101 | id^B = `lam (`emb (`var z)) 102 | \end{code} 103 | % 104 | \begin{code} 105 | where open PATTERNS 106 | \end{code} 107 | -------------------------------------------------------------------------------- /doc/generic-syntax.agda/Generic/Semantics.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | {-# OPTIONS --safe --sized-types #-} 3 | 4 | module Generic.Semantics where 5 | 6 | open import Size 7 | open import Data.List.Base as L hiding (lookup ; [_]) 8 | 9 | open import Data.Var hiding (z; s) 10 | open import Data.Var.Varlike using (VarLike; base) 11 | open import Data.Relation 12 | open import Relation.Unary 13 | open import Data.Environment 14 | open import Function using (flip) 15 | open import Generic.Syntax 16 | 17 | private 18 | variable 19 | I : Set 20 | σ : I 21 | Γ Δ : List I 22 | s : Size 23 | d : Desc I 24 | 25 | module _ {d : Desc I} where 26 | 27 | \end{code} 28 | %<*comp> 29 | \begin{code} 30 | _─Comp : List I → I ─Scoped → List I → Set 31 | (Γ ─Comp) 𝓒 Δ = ∀ {s σ} → Tm d s σ Γ → 𝓒 σ Δ 32 | \end{code} 33 | % 34 | \begin{code} 35 | private 36 | module DISPLAYONLY where 37 | \end{code} 38 | %<*semantics> 39 | \begin{code} 40 | record Semantics (d : Desc I) (𝓥 𝓒 : I ─Scoped) : Set where 41 | field th^𝓥 : Thinnable (𝓥 σ) 42 | var : ∀[ 𝓥 σ ⇒ 𝓒 σ ] 43 | alg : ∀[ ⟦ d ⟧ (Kripke 𝓥 𝓒) σ ⇒ 𝓒 σ ] 44 | \end{code} 45 | % 46 | %<*semrec> 47 | \begin{code} 48 | record Semantics (d : Desc I) (𝓥 𝓒 : I ─Scoped) : Set where 49 | \end{code} 50 | % 51 | \begin{code} 52 | field 53 | \end{code} 54 | %<*thv> 55 | \begin{code} 56 | th^𝓥 : Thinnable (𝓥 σ) 57 | \end{code} 58 | % 59 | %<*var> 60 | \begin{code} 61 | var : ∀[ 𝓥 σ ⇒ 𝓒 σ ] 62 | \end{code} 63 | % 64 | %<*alg> 65 | \begin{code} 66 | alg : ∀[ ⟦ d ⟧ (Kripke 𝓥 𝓒) σ ⇒ 𝓒 σ ] 67 | \end{code} 68 | % 69 | %<*semtype> 70 | %<*semanticstype> 71 | \begin{code} 72 | semantics : (Γ ─Env) 𝓥 Δ → (Γ ─Comp) 𝓒 Δ 73 | \end{code} 74 | % 75 | \begin{code} 76 | body : (Γ ─Env) 𝓥 Δ → ∀ Θ σ → 77 | Scope (Tm d s) Θ σ Γ → Kripke 𝓥 𝓒 Θ σ Δ 78 | \end{code} 79 | % 80 | %<*semproof> 81 | \begin{code} 82 | semantics ρ (`var k) = var (lookup ρ k) 83 | semantics ρ (`con t) = alg (fmap d (body ρ) t) 84 | \end{code} 85 | % 86 | %<*bodyproof> 87 | \begin{code} 88 | body ρ [] i t = semantics ρ t 89 | body ρ (_ ∷ _) i t = λ σ vs → semantics (vs ++^Env th^Env th^𝓥 ρ σ) t 90 | \end{code} 91 | % 92 | \begin{code} 93 | ◇-sem : (Γ ─◇Env) 𝓥 Δ → (Γ ─Comp) 𝓒 Δ 94 | ◇-body : (Γ ─◇Env) 𝓥 Δ → ∀ Θ σ → 95 | Scope (Tm d s) Θ σ Γ → Kripke 𝓥 𝓒 Θ σ Δ 96 | 97 | ◇-sem ρ (`var k) = var (DI.run th^𝓥 (slookup ρ k)) 98 | ◇-sem ρ (`con t) = alg (fmap d (◇-body ρ) t) 99 | 100 | ◇-body ρ [] i t = ◇-sem ρ t 101 | ◇-body ρ Δ@(_ ∷ _) i t = λ σ vs → ◇-sem (Δ ⊣ vs ,, ρ ◃ σ) t 102 | \end{code} 103 | %<*closed> 104 | \begin{code} 105 | closed : TM d σ → 𝓒 σ [] 106 | closed = semantics ε 107 | \end{code} 108 | % 109 | %<*eval> 110 | \begin{code} 111 | eval : VarLike 𝓥 → ∀[ Tm d s σ ⇒ 𝓒 σ ] 112 | eval vl^𝓥 = semantics (base vl^𝓥) 113 | \end{code} 114 | % 115 | -------------------------------------------------------------------------------- /doc/generic-syntax.agda/Generic/Semantics/Syntactic.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | {-# OPTIONS --safe --sized-types #-} 3 | 4 | open import Generic.Syntax using (Desc) 5 | 6 | module Generic.Semantics.Syntactic {I} {d : Desc I} where 7 | 8 | open import Size 9 | open import Data.List hiding ([_] ; lookup) 10 | open import Function 11 | open import Relation.Binary.PropositionalEquality hiding ([_]) 12 | open ≡-Reasoning 13 | 14 | open import Relation.Unary 15 | open import Data.Var 16 | open import Data.Var.Varlike 17 | open import Data.Environment 18 | open import Data.Relation 19 | open import Generic.Syntax 20 | open import Generic.Semantics 21 | 22 | open Semantics 23 | 24 | private 25 | variable 26 | σ τ : I 27 | Γ Δ : List I 28 | 29 | \end{code} 30 | %<*renaming> 31 | \begin{code} 32 | Ren : Semantics d Var (Tm d ∞) 33 | Ren .th^𝓥 = th^Var 34 | Ren .var = `var 35 | Ren .alg = `con ∘ fmap d (reify vl^Var) 36 | \end{code} 37 | % 38 | %<*thTm> 39 | \begin{code} 40 | th^Tm : Thinnable (Tm d ∞ σ) 41 | th^Tm t ρ = Semantics.semantics Ren ρ t 42 | \end{code} 43 | % 44 | \begin{code} 45 | 46 | vl^Tm : VarLike (Tm d ∞) 47 | new vl^Tm = `var z 48 | th^𝓥 vl^Tm = th^Tm 49 | \end{code} 50 | %<*substitution> 51 | \begin{code} 52 | Sub : Semantics d (Tm d ∞) (Tm d ∞) 53 | Sub .th^𝓥 = th^Tm 54 | Sub .var = id 55 | Sub .alg = `con ∘ fmap d (reify vl^Tm) 56 | \end{code} 57 | % 58 | \begin{code} 59 | module PAPERONLY where 60 | \end{code} 61 | %<*ren> 62 | \begin{code} 63 | ren : (Γ ─Env) Var Δ → 64 | Tm d ∞ σ Γ → Tm d ∞ σ Δ 65 | ren ρ t = Semantics.semantics Ren ρ t 66 | \end{code} 67 | % 68 | %<*sub> 69 | \begin{code} 70 | sub : (Γ ─Env) (Tm d ∞) Δ → 71 | Tm d ∞ σ Γ → Tm d ∞ σ Δ 72 | sub ρ t = Semantics.semantics Sub ρ t 73 | \end{code} 74 | % 75 | \begin{code} 76 | ren : Thinning Γ Δ → (Γ ─Comp) (Tm d ∞) Δ 77 | ren = Semantics.semantics Ren 78 | 79 | sub : ∀ {s} → (Γ ─Env) (Tm d ∞) Δ → Tm d s σ Γ → Tm d ∞ σ Δ 80 | sub ρ t = Semantics.semantics Sub ρ t 81 | 82 | vl^VarTm : VarLikeᴿ VarTmᴿ vl^Var vl^Tm 83 | VarLikeᴿ.newᴿ vl^VarTm = refl 84 | VarLikeᴿ.thᴿ vl^VarTm = λ σ → cong (ren σ) 85 | 86 | reify^Tm : ∀ Δ → ∀[ Kripke (Tm d ∞) (Tm d ∞) Δ σ ⇒ (Δ ++_) ⊢ Tm d ∞ σ ] 87 | reify^Tm Δ = reify vl^Tm Δ _ 88 | 89 | id^Tm : (Γ ─Env) (Tm d ∞) Γ 90 | lookup id^Tm = `var 91 | 92 | lookup-base^Tm : (k : Var σ Γ) → lookup (base vl^Tm) k ≡ `var k 93 | lookup-base^Tm z = refl 94 | lookup-base^Tm (s k) rewrite lookup-base^Tm k = refl 95 | 96 | base^VarTmᴿ : ∀ {Γ} → All VarTmᴿ Γ (base vl^Var) (base vl^Tm) 97 | lookupᴿ base^VarTmᴿ k = begin 98 | `var (lookup (base vl^Var) k) ≡⟨ cong `var (lookup-base^Var k) ⟩ 99 | `var k ≡⟨ sym (lookup-base^Tm k) ⟩ 100 | lookup (base vl^Tm) k ∎ 101 | 102 | infix 5 _[_ 103 | infix 6 _/0] 104 | 105 | _/0] : Tm d ∞ σ Γ → (σ ∷ Γ ─Env) (Tm d ∞) Γ 106 | _/0] = singleton vl^Tm 107 | 108 | _[_ : Tm d ∞ τ (σ ∷ Γ) → (σ ∷ Γ ─Env) (Tm d ∞) Γ → Tm d ∞ τ Γ 109 | t [ ρ = sub ρ t 110 | \end{code} 111 | -------------------------------------------------------------------------------- /doc/generic-syntax.agda/Generic/Semantics/TypeChecking.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | {-# OPTIONS --safe --sized-types #-} 3 | 4 | module Generic.Semantics.TypeChecking where 5 | 6 | open import Size 7 | open import Function 8 | open import Data.Unit using (⊤; tt) 9 | open import Data.Product 10 | open import Data.List hiding ([_]) 11 | open import Data.Maybe using (Maybe; nothing; just) 12 | import Data.Maybe.Categorical as MC 13 | 14 | open import Data.Var hiding (_<$>_) 15 | open import Data.Environment hiding (_<$>_) 16 | open import Generic.Syntax 17 | open import Generic.Syntax.Bidirectional; open PATTERNS 18 | open import Generic.Semantics 19 | 20 | import Category.Monad as CM 21 | import Level 22 | module M = CM.RawMonad (MC.monad {Level.zero}) 23 | open M 24 | 25 | open import Relation.Binary.PropositionalEquality hiding ([_]) 26 | 27 | infix 2 _=?_ 28 | \end{code} 29 | %<*typeeq> 30 | \begin{code} 31 | _=?_ : (σ τ : Type) → Maybe ⊤ 32 | α =? α = just tt 33 | (σ `→ τ) =? (φ `→ ψ) = (σ =? φ) >> (τ =? ψ) 34 | _ =? _ = nothing 35 | \end{code} 36 | % 37 | %<*isArrow> 38 | \begin{code} 39 | isArrow : Type → Maybe (Type × Type) 40 | isArrow (σ `→ τ) = just (σ , τ) 41 | isArrow _ = nothing 42 | \end{code} 43 | % 44 | 45 | \begin{code} 46 | 47 | private 48 | variable 49 | i : Mode 50 | Γ : List Mode 51 | 52 | \end{code} 53 | 54 | %<*typemode> 55 | \begin{code} 56 | Type- : Mode → Set 57 | Type- Check = Type → Maybe ⊤ 58 | Type- Synth = Maybe Type 59 | \end{code} 60 | % 61 | %<*varmode> 62 | \begin{code} 63 | data Var- : Mode → Set where 64 | `var : Type → Var- Synth 65 | \end{code} 66 | % 67 | 68 | %<*app> 69 | \begin{code} 70 | app : Type- Synth → Type- Check → Type- Synth 71 | app f t = do 72 | arr ← f 73 | (σ , τ) ← isArrow arr 74 | τ <$ t σ 75 | \end{code} 76 | % 77 | %<*cut> 78 | \begin{code} 79 | cut : Type → Type- Check → Type- Synth 80 | cut σ t = σ <$ t σ 81 | \end{code} 82 | % 83 | %<*emb> 84 | \begin{code} 85 | emb : Type- Synth → Type- Check 86 | emb t σ = do 87 | τ ← t 88 | σ =? τ 89 | \end{code} 90 | % 91 | %<*lam> 92 | \begin{code} 93 | lam : Kripke (const ∘ Var-) (const ∘ Type-) (Synth ∷ []) Check Γ → Type- Check 94 | lam b arr = do 95 | (σ , τ) ← isArrow arr 96 | b (bind Synth) (ε ∙ `var σ) τ 97 | \end{code} 98 | % 99 | 100 | %<*typecheck> 101 | \begin{code} 102 | Typecheck : Semantics Bidi (const ∘ Var-) (const ∘ Type-) 103 | Semantics.th^𝓥 Typecheck = th^const 104 | Semantics.var Typecheck = λ where (`var σ) → just σ 105 | Semantics.alg Typecheck = λ where 106 | (`app' f t) → app f t 107 | (`cut' σ t) → cut σ t 108 | (`emb' t) → emb t 109 | (`lam' b) → lam b 110 | \end{code} 111 | % 112 | 113 | %<*type-> 114 | \begin{code} 115 | type- : ∀ mode → TM Bidi mode → Type- mode 116 | type- p = Semantics.closed Typecheck 117 | \end{code} 118 | % 119 | \begin{code} 120 | 121 | \end{code} 122 | %<*example> 123 | \begin{code} 124 | _ : let β = α `→ α in type- Synth (`app (`cut (β `→ β) id^B) id^B) ≡ just β 125 | _ = refl 126 | \end{code} 127 | % 128 | -------------------------------------------------------------------------------- /doc/shared/Data/Relation.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | {-# OPTIONS --safe --sized-types #-} 3 | 4 | module Data.Relation where 5 | 6 | open import Size 7 | open import Data.Sum 8 | open import Data.List.Base hiding (lookup ; [_]) 9 | 10 | open import Data.Var hiding (_<$>_) 11 | open import Data.Environment 12 | open import Generic.Syntax 13 | open import Relation.Unary hiding (U) 14 | open import Agda.Builtin.Equality 15 | open import Function 16 | 17 | private 18 | variable 19 | I : Set 20 | σ : I 21 | T U : I ─Scoped 22 | 𝓥ᴬ 𝓥ᴮ : I ─Scoped 23 | Γ Δ : List I 24 | 25 | \end{code} 26 | %<*rel> 27 | \begin{code} 28 | record Rel (T U : I ─Scoped) : Set₁ where 29 | constructor mkRel 30 | field rel : ∀ σ → ∀[ T σ ⇒ U σ ⇒ const Set ] 31 | \end{code} 32 | % 33 | \begin{code} 34 | open Rel public 35 | 36 | \end{code} 37 | %<*all> 38 | \begin{code} 39 | record All (𝓥ᴿ : Rel 𝓥ᴬ 𝓥ᴮ) (Γ : List I) 40 | (ρᴬ : (Γ ─Env) 𝓥ᴬ Δ) (ρᴮ : (Γ ─Env) 𝓥ᴮ Δ) : Set where 41 | constructor packᴿ 42 | field lookupᴿ : ∀ k → rel 𝓥ᴿ σ (lookup ρᴬ k) (lookup ρᴮ k) 43 | \end{code} 44 | % 45 | \begin{code} 46 | open All public 47 | 48 | module _ {T U : I ─Scoped} {𝓡 : Rel T U} where 49 | 50 | private 51 | variable 52 | ρᵀ σᵀ : (Γ ─Env) T Δ 53 | ρᵁ σᵁ : (Γ ─Env) U Δ 54 | vᵀ : T σ Γ 55 | vᵁ : U σ Γ 56 | fᵀ : ∀ {i} → T i Γ → T i Δ 57 | fᵁ : ∀ {i} → U i Γ → U i Δ 58 | 59 | εᴿ : All 𝓡 [] ρᵀ ρᵁ 60 | lookupᴿ εᴿ () 61 | 62 | infixl 20 _∙ᴿ_ _∷ᴿ_ 63 | _∙ᴿ_ : All 𝓡 Γ ρᵀ ρᵁ → rel 𝓡 σ vᵀ vᵁ → All 𝓡 (σ ∷ Γ) (ρᵀ ∙ vᵀ) (ρᵁ ∙ vᵁ) 64 | lookupᴿ (ρ ∙ᴿ v) z = v 65 | lookupᴿ (ρ ∙ᴿ v) (s k) = lookupᴿ ρ k 66 | 67 | _∷ᴿ_ : rel 𝓡 σ (lookup ρᵀ z) (lookup ρᵁ z) → 68 | (∀ {σ} (v : Var σ Γ) → rel 𝓡 _ (lookup ρᵀ (s v)) (lookup ρᵁ (s v))) → 69 | All 𝓡 (σ ∷ Γ) ρᵀ ρᵁ 70 | lookupᴿ (v ∷ᴿ ρ) z = v 71 | lookupᴿ (v ∷ᴿ ρ) (s k) = ρ k 72 | 73 | _++^Envᴿ_ : All 𝓡 Γ ρᵀ ρᵁ → All 𝓡 Δ σᵀ σᵁ → 74 | All 𝓡 (Γ ++ Δ) (ρᵀ ++^Env σᵀ) (ρᵁ ++^Env σᵁ) 75 | lookupᴿ (_++^Envᴿ_ {Γ} ρᴿ σᴿ) k with split Γ k 76 | ... | inj₁ k₁ = lookupᴿ ρᴿ k₁ 77 | ... | inj₂ k₂ = lookupᴿ σᴿ k₂ 78 | 79 | selectᴿ : ∀ ρ → All 𝓡 Δ ρᵀ ρᵁ → All 𝓡 Γ (select ρ ρᵀ) (select ρ ρᵁ) 80 | lookupᴿ (selectᴿ ρ ρᴿ) k = lookupᴿ ρᴿ (lookup ρ k) 81 | 82 | _<$>ᴿ_ : (∀ {i t u} → rel 𝓡 i t u → rel 𝓡 i (fᵀ t) (fᵁ u)) → 83 | All 𝓡 Γ ρᵀ ρᵁ → All 𝓡 Γ (fᵀ <$> ρᵀ) (fᵁ <$> ρᵁ) 84 | lookupᴿ (F <$>ᴿ ρ) k = F (lookupᴿ ρ k) 85 | 86 | module _ {T : I ─Scoped} where 87 | 88 | private 89 | variable 90 | ρ : (Γ ─Env) T Δ 91 | 92 | \end{code} 93 | %<*eqR> 94 | \begin{code} 95 | Eqᴿ : Rel T T 96 | rel Eqᴿ i = _≡_ 97 | \end{code} 98 | % 99 | %<*reflR> 100 | \begin{code} 101 | reflᴿ : All Eqᴿ Γ ρ ρ 102 | lookupᴿ reflᴿ k = refl 103 | \end{code} 104 | % 105 | \begin{code} 106 | 107 | module _ {A B : I ─Scoped} where 108 | 109 | open import Relation.Binary.HeterogeneousEquality.Core 110 | 111 | HEqᴿ : Rel A B 112 | rel HEqᴿ i = λ a b → a ≅ b 113 | 114 | module _ {d : Desc I} where 115 | 116 | VarTmᴿ : Rel Var (Tm d ∞) 117 | rel VarTmᴿ i v t = `var v ≡ t 118 | \end{code} 119 | -------------------------------------------------------------------------------- /doc/generic-syntax.agda/Generic/Scopecheck.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | {-# OPTIONS --safe --sized-types #-} 3 | 4 | open import Relation.Binary using (Decidable) 5 | open import Relation.Binary.PropositionalEquality 6 | 7 | open import Generic.Syntax 8 | 9 | module Generic.Scopecheck {E I : Set} (_≟I_ : Decidable {A = I} _≡_) where 10 | 11 | open import Category.Monad 12 | 13 | open import Level 14 | open import Size 15 | open import Data.List.Base using (List; []; _∷_) 16 | open import Data.List.Relation.Unary.All using (All; []; _∷_) 17 | open import Data.List.Relation.Unary.All.Properties 18 | using () renaming (++⁺ to _++_) 19 | 20 | open import Data.Product 21 | open import Data.String using (String; _≟_) 22 | open import Data.Sum 23 | import Data.Sum.Categorical.Left as SC 24 | open import Function 25 | open import Relation.Nullary 26 | 27 | open import Data.Var hiding (_<$>_) 28 | 29 | private 30 | variable 31 | σ : I 32 | Γ Δ : List I 33 | i : Size 34 | A : Set 35 | 36 | \end{code} 37 | %<*names> 38 | \begin{code} 39 | Names : List I → Set 40 | Names = All (const String) 41 | \end{code} 42 | % 43 | %<*withnames> 44 | \begin{code} 45 | WithNames : (I → Set) → List I → I ─Scoped 46 | WithNames T [] j Γ = T j 47 | WithNames T Δ j Γ = Names Δ × T j 48 | \end{code} 49 | % 50 | %<*raw> 51 | \begin{code} 52 | data Raw (d : Desc I) : Size → I → Set where 53 | `var : E → String → Raw d (↑ i) σ 54 | `con : ⟦ d ⟧ (WithNames (Raw d i)) σ [] → Raw d (↑ i) σ 55 | \end{code} 56 | % 57 | \begin{code} 58 | 59 | \end{code} 60 | %<*error> 61 | \begin{code} 62 | data Error : Set where 63 | OutOfScope : Error 64 | WrongSort : (σ τ : I) → σ ≢ τ → Error 65 | \end{code} 66 | % 67 | \begin{code} 68 | 69 | private 70 | \end{code} 71 | %<*monad> 72 | \begin{code} 73 | Fail : Set → Set 74 | Fail A = (Error × E × String) ⊎ A 75 | \end{code} 76 | % 77 | %<*fail> 78 | \begin{code} 79 | fail : Error → E → String → Fail A 80 | fail err e str = inj₁ (err , e , str) 81 | \end{code} 82 | % 83 | \begin{code} 84 | open RawMonad (SC.monad (Error × E × String) zero) 85 | 86 | instance _ = rawIApplicative 87 | 88 | \end{code} 89 | %<*toVar> 90 | \begin{code} 91 | toVar : E → String → ∀ σ Γ → Names Γ → Fail (Var σ Γ) 92 | toVar e x σ [] [] = fail OutOfScope e x 93 | toVar e x σ (τ ∷ Γ) (y ∷ scp) with x ≟ y | σ ≟I τ 94 | ... | yes _ | yes refl = pure z 95 | ... | yes _ | no ¬eq = fail (WrongSort σ τ ¬eq) e x 96 | ... | no ¬p | _ = s <$> toVar e x σ Γ scp 97 | \end{code} 98 | % 99 | \begin{code} 100 | module _ {d : Desc I} where 101 | \end{code} 102 | %<*scopecheck> 103 | \begin{AgdaAlign} 104 | \begin{AgdaSuppressSpace} 105 | %<*totmtype> 106 | \begin{code} 107 | toTm : Names Γ → Raw d i σ → Fail (Tm d i σ Γ) 108 | \end{code} 109 | % 110 | \begin{code} 111 | toScope : Names Γ → ∀ Δ σ → WithNames (Raw d i) Δ σ [] → Fail (Scope (Tm d i) Δ σ Γ) 112 | 113 | toTm scp (`var e v) = `var <$> toVar e v _ _ scp 114 | toTm scp (`con b) = `con <$> mapA d (toScope scp) b 115 | 116 | toScope scp [] σ b = toTm scp b 117 | toScope scp Δ@(_ ∷ _) σ (bnd , b) = toTm (bnd ++ scp) b 118 | \end{code} 119 | \end{AgdaSuppressSpace} 120 | \end{AgdaAlign} 121 | % 122 | -------------------------------------------------------------------------------- /doc/titlingpage.tex: -------------------------------------------------------------------------------- 1 | \begin{titlingpage} 2 | \begin{center} 3 | \vspace*{1cm} 4 | 5 | \Huge 6 | \textbf{Syntaxes with Binding, \\ Their Programs, and Proofs} 7 | 8 | \vspace{1.5cm} 9 | 10 | \LARGE 11 | 12 | \textbf{Guillaume Xavier ALLAIS} 13 | 14 | \vfill 15 | 16 | A thesis presented for the degree of\\ 17 | Doctor of Philosophy 18 | 19 | \vspace{0.8cm} 20 | 21 | \includegraphics[width=0.4\textwidth]{strath_coat} 22 | 23 | \vspace{0.8cm} 24 | 25 | \Large 26 | Computer and Information Sciences\\ 27 | University of Strathclyde\\ 28 | United Kingdom\\ 29 | 2022 30 | 31 | \end{center} 32 | \end{titlingpage} 33 | 34 | 35 | \pagebreak 36 | { 37 | \hspace{0pt} 38 | \vfill{} 39 | This thesis is the result of the author’s original research. It has been 40 | composed by the author and has not been previously submitted for 41 | examination which has led to the award of a degree. 42 | 43 | The copyright of this thesis belongs to the author under the terms of the 44 | United Kingdom Copyright Acts as qualified by University of Strathclyde 45 | Regulation 3.50. Due acknowledgement must always be made of the use of 46 | any material contained in, or derived from, this thesis. 47 | 48 | \vspace{1.5cm} 49 | 50 | \hfill 51 | \begin{minipage}{0.4\textwidth} 52 | Signed: Guillaume ALLAIS \\ 53 | Date: \today{} 54 | \end{minipage} 55 | \vfill 56 | \hspace{0pt} 57 | } 58 | \pagebreak 59 | 60 | \input{acknowledgments} 61 | \pagebreak 62 | 63 | \hspace{0pt} 64 | \vfill{} 65 | \begin{abstract} 66 | Almost every programming language's syntax includes a notion of binder and 67 | corresponding bound occurrences, along with the accompanying notions of 68 | α-equivalence, capture-avoiding substitution, typing contexts, 69 | runtime environments, and so on. 70 | 71 | In the past, implementing and reasoning about programming languages required 72 | careful handling to maintain the correct behaviour of bound variables. Modern 73 | programming languages include features that enable constraints like scope 74 | safety to be expressed in types. 75 | % 76 | Nevertheless, the programmer is still forced to write the same boilerplate 77 | over again for each new implementation of a scope safe operation (e.g., 78 | renaming, substitution, desugaring, printing, etc.), and then again 79 | for correctness proofs. 80 | 81 | In a case study focusing on the simply typed lambda calculus, we analyse these 82 | well scoped traversals and observe that they all share the same structure. 83 | % 84 | This enables us to formulate them as instances of a more general program whose 85 | properties can be established generically. 86 | % 87 | Alas, the programmer is still forced to duplicate this effort for every new 88 | language they implement. 89 | 90 | This leads us to defining an expressive universe of syntaxes with binding 91 | and to demonstrate how to implement scope safe traversals once and for all by 92 | generic programming and how to derive properties of these traversals by 93 | generic proving. 94 | % 95 | Our universe description, generic traversals and proofs, and our examples have 96 | all been formalised in Agda. 97 | \end{abstract} 98 | \vfill 99 | \hspace{0pt} 100 | \pagebreak 101 | -------------------------------------------------------------------------------- /doc/type-scope-semantics.agda/Semantics/Specification.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | {-# OPTIONS --safe --sized-types #-} 3 | 4 | module Semantics.Specification where 5 | 6 | open import Syntax.Type 7 | open import Syntax.Calculus 8 | open import Data.Var 9 | open import Data.Environment hiding (Kripke) 10 | open import Data.List using (List; []; _∷_) 11 | open import Relation.Unary 12 | 13 | private 14 | 15 | variable 16 | σ τ : Type 17 | Γ Δ : List Type 18 | 𝓥 𝓒 : Type ─Scoped 19 | 20 | \end{code} 21 | %<*comp> 22 | \begin{code} 23 | _─Comp : List Type → Type ─Scoped → List Type → Set 24 | (Γ ─Comp) 𝓒 Δ = ∀ {σ} → Term σ Γ → 𝓒 σ Δ 25 | \end{code} 26 | % 27 | 28 | %<*kripke> 29 | \begin{code} 30 | Kripke : (𝓥 𝓒 : Type ─Scoped) → Type → Type → List Type → Set 31 | Kripke 𝓥 𝓒 σ τ Γ = ∀[ Thinning Γ ⇒ 𝓥 σ ⇒ 𝓒 τ ] 32 | \end{code} 33 | % 34 | %<*wholerecsem> 35 | %<*recsem> 36 | \begin{code} 37 | record Semantics (𝓥 𝓒 : Type ─Scoped) : Set where 38 | \end{code} 39 | % 40 | \begin{code} 41 | field 42 | \end{code} 43 | %<*recfields> 44 | %<*thV> 45 | \begin{code} 46 | th^𝓥 : Thinnable (𝓥 σ) 47 | \end{code} 48 | % 49 | %<*var> 50 | \begin{code} 51 | var : ∀[ 𝓥 σ ⇒ 𝓒 σ ] 52 | \end{code} 53 | % 54 | %<*lam> 55 | \begin{code} 56 | lam : ∀[ □ (𝓥 σ ⇒ 𝓒 τ) ⇒ 𝓒 (σ `→ τ) ] 57 | \end{code} 58 | % 59 | %<*app> 60 | \begin{code} 61 | app : ∀[ 𝓒 (σ `→ τ) ⇒ 𝓒 σ ⇒ 𝓒 τ ] 62 | \end{code} 63 | % 64 | %<*one> 65 | \begin{code} 66 | one : ∀[ 𝓒 `Unit ] 67 | \end{code} 68 | % 69 | %<*bool> 70 | \begin{code} 71 | tt : ∀[ 𝓒 `Bool ] 72 | ff : ∀[ 𝓒 `Bool ] 73 | \end{code} 74 | % 75 | %<*ifte> 76 | \begin{code} 77 | ifte : ∀[ 𝓒 `Bool ⇒ 𝓒 σ ⇒ 𝓒 σ ⇒ 𝓒 σ ] 78 | \end{code} 79 | % 80 | % 81 | % 82 | \begin{code} 83 | Evaluation : (𝓥 𝓒 : Type ─Scoped) → Set 84 | Evaluation 𝓥 𝓒 = ∀ {Γ Δ} → (Γ ─Env) 𝓥 Δ → (Γ ─Comp) 𝓒 Δ 85 | 86 | Evaluation' : (𝓒 : Type ─Scoped) → Set 87 | Evaluation' 𝓒 = ∀ {Γ} → (Γ ─Comp) 𝓒 Γ 88 | 89 | \end{code} 90 | %<*fundamental> 91 | %<*fundamentalheader> 92 | \begin{code} 93 | module _ (𝓢 : Semantics 𝓥 𝓒) where 94 | open Semantics 𝓢 95 | \end{code} 96 | % 97 | %<*fundamentalbody> 98 | %<*semantics-type> 99 | \begin{code} 100 | semantics : (Γ ─Env) 𝓥 Δ → (Γ ─Comp) 𝓒 Δ 101 | \end{code} 102 | % 103 | %<*semantics-var> 104 | \begin{code} 105 | semantics ρ (`var v) = var (lookup ρ v) 106 | \end{code} 107 | % 108 | %<*semantics-app> 109 | \begin{code} 110 | semantics ρ (`app t u) = app (semantics ρ t) (semantics ρ u) 111 | \end{code} 112 | % 113 | %<*semantics-lam> 114 | \begin{code} 115 | semantics ρ (`lam t) = lam (λ re u → semantics (th^Env th^𝓥 ρ re ∙ u) t) 116 | \end{code} 117 | % 118 | %<*semantics-one> 119 | \begin{code} 120 | semantics ρ `one = one 121 | \end{code} 122 | % 123 | %<*semantics-bool> 124 | \begin{code} 125 | semantics ρ `tt = tt 126 | semantics ρ `ff = ff 127 | \end{code} 128 | % 129 | %<*semantics-ifte> 130 | \begin{code} 131 | semantics ρ (`ifte b l r) = ifte (semantics ρ b) (semantics ρ l) (semantics ρ r) 132 | \end{code} 133 | % 134 | % 135 | % 136 | -------------------------------------------------------------------------------- /doc/thesis.tex: -------------------------------------------------------------------------------- 1 | \documentclass{memoir} 2 | 3 | \usepackage{tocbasic} 4 | \DeclareTOCStyleEntry[dynnumwidth]{tocline}{figure}% for figure entries 5 | \DeclareTOCStyleEntry[dynnumwidth]{tocline}{table} 6 | 7 | \usepackage{agda} 8 | \usepackage{catchfilebetweentags} 9 | \usepackage[T1]{fontenc} 10 | \usepackage{mathpartir} 11 | \usepackage{natbib} 12 | \usepackage{todonotes} 13 | 14 | \usepackage{hyperref} 15 | \usepackage{cleveref} 16 | \usepackage{afterpage} 17 | \usepackage{tcolorbox} 18 | %\usepackage{savetrees} 19 | %\usepackage{showframe} 20 | \usepackage{microtype} 21 | 22 | \usepackage[pass]{geometry} 23 | 24 | %% disable the >> ligatures 25 | \DisableLigatures[>]{} 26 | %% disable the -- ligatures 27 | \DisableLigatures[-]{} 28 | 29 | %\renewcommand{\baselinestretch}{1.5} 30 | \bibliographystyle{abbrvnat} 31 | 32 | \setlength\mathindent{0.5em} 33 | 34 | \usepackage{thmtools} 35 | \usepackage{thm-restate} 36 | \usepackage{xparse} 37 | \input{convention} 38 | 39 | 40 | \newtheorem{definition}{Definition} 41 | \newtheorem{lemma}{Lemma} 42 | \newtheorem{theorem}{Theorem} 43 | \newtheorem{corollary}{Corollary} 44 | \newtheorem{example}{Example} 45 | 46 | %\newtheorem{convention}{Convention} 47 | %\newtheorem{feature}{Feature} 48 | %\newtheorem{technique}{Technique} 49 | \newtheorem{remark}{Remark} 50 | 51 | \newenvironment{agdasnippet}[1][0] 52 | {\small\ifx{#1}\undefined 53 | \centering 54 | \else 55 | \fi 56 | }{\ignorespacesafterend% 57 | } 58 | 59 | \input{unicode} 60 | \input{commands} 61 | \input{robust-catch} 62 | 63 | \title{\Huge Syntaxes with Binding, \\ Their Programs, and Proofs} 64 | \author{\Large Guillaume ALLAIS} 65 | 66 | \begin{document} 67 | 68 | \frontmatter 69 | \input{cover} 70 | \cleardoublepage{} 71 | 72 | \maketitle{} 73 | 74 | \mainmatter 75 | \pagestyle{ruled} 76 | \begin{KeepFromToc} 77 | \tableofcontents 78 | \end{KeepFromToc} 79 | 80 | \input{introduction.tex/intro} 81 | %\input{introduction.tex/type-theory} 82 | \input{introduction.tex/category} 83 | \input{introduction.tex/agda} 84 | 85 | 86 | \part{Type and Scope Preserving Programs} 87 | \label{type-scope-semantics} 88 | \input{type-scope-semantics.tex/syntax.tex} 89 | \input{type-scope-semantics.tex/refactoring.tex} 90 | \input{type-scope-semantics.tex/normalization.tex} 91 | \input{type-scope-semantics.tex/cps-translation.tex} 92 | \input{type-scope-semantics.tex/program-conclusion.tex} 93 | 94 | \part{And Their Proofs} 95 | \label{properties} 96 | \input{type-scope-semantics.tex/simulation.tex} 97 | \input{type-scope-semantics.tex/fusion.tex} 98 | 99 | \input{type-scope-semantics.tex/proof-conclusion.tex} 100 | 101 | \part{A Universe of Well Kinded-and-Scoped Syntaxes with Binding, their Programs and their Proofs} 102 | \label{a-universe} 103 | 104 | \input{generic-syntax.tex/a-universe.tex} 105 | \input{generic-syntax.tex/data.tex} 106 | \input{generic-syntax.tex/syntax.tex} 107 | 108 | \input{generic-syntax.tex/safe-programs.tex} 109 | \input{generic-syntax.tex/compiling.tex} 110 | \input{generic-syntax.tex/other-programs.tex} 111 | 112 | \input{generic-syntax.tex/proofs.tex} 113 | 114 | \input{conclusion} 115 | 116 | \cleardoublepage{} 117 | 118 | 119 | \appendix 120 | 121 | \chapter{Conventions and Techniques}\label{sec:conventions} 122 | \allconventions 123 | \alltechniques 124 | 125 | \chapter{Agda Features}\label{sec:features} 126 | \allfeatures 127 | 128 | \clearpage 129 | \listoffigures 130 | \bibliography{main} 131 | 132 | \end{document} 133 | -------------------------------------------------------------------------------- /doc/generic-syntax.agda/Generic/Semantics/Elaboration/LetCounter.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | {-# OPTIONS --safe --sized-types #-} 3 | 4 | open import Generic.Syntax 5 | 6 | module Generic.Semantics.Elaboration.LetCounter {I : Set} {d : Desc I} where 7 | 8 | import Level as L 9 | open import Size 10 | open import Agda.Builtin.Equality 11 | open import Agda.Builtin.Bool 12 | open import Data.Product 13 | import Data.Product.Categorical.Right as PC 14 | open import Data.List.Base using (List; []; _∷_) 15 | open import Data.List.Relation.Unary.All as All using (All; _∷_) 16 | open import Data.List.Relation.Unary.All.Properties renaming (++⁻ʳ to drop) 17 | open import Function 18 | 19 | open import Relation.Unary 20 | open import Data.Var 21 | open import Data.Var.Varlike 22 | open import Data.Environment using (Kripke; th^Var; ε; _∙_; identity; extend; extract) 23 | open import Generic.Syntax.LetCounter 24 | open import Generic.Syntax.LetBinder 25 | open import Generic.Semantics 26 | open import Generic.Semantics.Syntactic 27 | 28 | private 29 | variable 30 | Γ Δ : List I 31 | σ : I 32 | 33 | module PCR {Γ : List I} = PC L.zero (rawMonoid Γ) 34 | instance _ = PCR.applicative 35 | 36 | \end{code} 37 | %<*counted> 38 | \begin{code} 39 | Counted : I ─Scoped → I ─Scoped 40 | Counted T i Γ = T i Γ × Count Γ 41 | \end{code} 42 | % 43 | %<*reifycount> 44 | \begin{code} 45 | reify^Count : ∀ Δ σ → Kripke Var (Counted (Tm (d `+ CLet) ∞)) Δ σ Γ → 46 | Counted (Scope (Tm (d `+ CLet) ∞) Δ) σ Γ 47 | reify^Count Δ σ kr = let (scp , c) = reify vl^Var Δ σ kr in scp , drop Δ c 48 | \end{code} 49 | % 50 | %<*letcount> 51 | \begin{code} 52 | clet : ⟦ Let ⟧ (Kripke Var (Counted (Tm (d `+ CLet) ∞))) σ Γ → 53 | Counted (⟦ CLet ⟧ (Scope (Tm (d `+ CLet) ∞))) σ Γ 54 | clet (στ , (e , ce) , body , eq) = case body extend (ε ∙ z) of λ where 55 | (t , cx ∷ ct) → (cx , στ , e , t , eq) , merge (control cx ce) ct 56 | \end{code} 57 | % 58 | 59 | %<*Annotate> 60 | \begin{code} 61 | Annotate : Semantics (d `+ Let) Var (Counted (Tm (d `+ CLet) ∞)) 62 | Semantics.th^𝓥 Annotate = th^Var 63 | Semantics.var Annotate = λ v → `var v , fromVar v 64 | Semantics.alg Annotate = λ where 65 | (true , t) → let (t' , c) = mapA d reify^Count t in `con (true , t') , c 66 | (false , et) → let (et' , c) = clet et in `con (false , et') , c 67 | \end{code} 68 | % 69 | 70 | %<*annotate> 71 | \begin{AgdaAlign} 72 | \begin{AgdaSuppressSpace} 73 | %<*annotatetype> 74 | \begin{code} 75 | annotate : ∀[ Tm (d `+ Let) ∞ σ ⇒ Tm (d `+ CLet) ∞ σ ] 76 | \end{code} 77 | % 78 | \begin{code} 79 | annotate t = let (t' , _) = Semantics.semantics Annotate identity t in t' 80 | \end{code} 81 | \end{AgdaSuppressSpace} 82 | \end{AgdaAlign} 83 | % 84 | \begin{code} 85 | 86 | Inline : Semantics (d `+ CLet) (Tm (d `+ Let) ∞) (Tm (d `+ Let) ∞) 87 | Semantics.th^𝓥 Inline = th^Tm 88 | Semantics.var Inline = id 89 | Semantics.alg Inline = λ where 90 | (true , t) → `con (true , fmap d (reify vl^Tm) t) 91 | (false , many , στ , e , b , eq) → `con (false , στ , e , b extend (ε ∙ `var z) , eq) 92 | (false , _ , στ , e , b , refl) → extract b (ε ∙ e) -- cf Semantics.alg UnLet 93 | 94 | inline : Tm (d `+ CLet) ∞ σ Γ → Tm (d `+ Let) ∞ σ Γ 95 | inline = Semantics.semantics Inline id^Tm 96 | 97 | inline-affine : Tm (d `+ Let) ∞ σ Γ → Tm (d `+ Let) ∞ σ Γ 98 | inline-affine = inline ∘′ annotate 99 | \end{code} 100 | -------------------------------------------------------------------------------- /doc/type-scope-semantics.agda/Properties/Fusion/Syntactic/Specification.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | open import Syntax.Type 3 | open import Data.Var 4 | 5 | module Properties.Fusion.Syntactic.Specification {𝓣ᴬ 𝓣ᴮ 𝓣ᴬᴮ : Type ─Scoped} where 6 | 7 | open import Data.Environment 8 | open import Data.List.Base using (List; []; _∷_) 9 | open import Data.Relation hiding (_∙ᴿ_) 10 | open import Syntax.Type 11 | open import Syntax.Calculus 12 | open import Semantics.Specification 13 | open import Semantics.Syntactic.Specification as SynSpec 14 | open import Function renaming (_$′_ to _$_) using (id) 15 | open import Properties.Fusion.Specification 16 | open import Relation.Unary 17 | open import Relation.Binary.PropositionalEquality.Extra 18 | 19 | private 20 | variable 21 | σ τ : Type 22 | Γ Δ Θ Ω : List Type 23 | ρᴬ : (Γ ─Env) 𝓣ᴬ Δ 24 | ρᴮ : (Δ ─Env) 𝓣ᴮ Θ 25 | ρᴬᴮ : (Γ ─Env) 𝓣ᴬᴮ Θ 26 | tᴮ : 𝓣ᴮ σ Γ 27 | tᴬᴮ : 𝓣ᴬᴮ σ Γ 28 | \end{code} 29 | %<*synfusion> 30 | \begin{code} 31 | record SynFusion 32 | (Synᴬ : Syntactic 𝓣ᴬ) (Synᴮ : Syntactic 𝓣ᴮ) (Synᴬᴮ : Syntactic 𝓣ᴬᴮ) 33 | (𝓔ᴿ : ∀ {Γ Δ Θ} → (Γ ─Env) 𝓣ᴬ Δ → (Δ ─Env) 𝓣ᴮ Θ → (Γ ─Env) 𝓣ᴬᴮ Θ → Set) 34 | (𝓣ᴿ : Rel 𝓣ᴮ 𝓣ᴬᴮ) : Set where 35 | \end{code} 36 | % 37 | \begin{code} 38 | module Synᴬ = Syntactic Synᴬ 39 | module Synᴮ = Syntactic Synᴮ 40 | module Synᴬᴮ = Syntactic Synᴬᴮ 41 | evalᴬ = Semantics.Specification.semantics (SynSpec.syntactic Synᴬ) 42 | evalᴮ = Semantics.Specification.semantics (SynSpec.syntactic Synᴮ) 43 | evalᴬᴮ = Semantics.Specification.semantics (SynSpec.syntactic Synᴬᴮ) 44 | \end{code} 45 | %<*crel> 46 | \begin{code} 47 | 𝓡 : ∀ σ → (Γ ─Env) 𝓣ᴬ Δ → (Δ ─Env) 𝓣ᴮ Θ → (Γ ─Env) 𝓣ᴬᴮ Θ → 48 | Term σ Γ → Set 49 | 𝓡 σ ρᴬ ρᴮ ρᴬᴮ t = evalᴮ ρᴮ (evalᴬ ρᴬ t) ≡ evalᴬᴮ ρᴬᴮ t 50 | \end{code} 51 | % 52 | \begin{code} 53 | field 54 | \end{code} 55 | %<*thV> 56 | \begin{code} 57 | _∙ᴿ_ : 𝓔ᴿ ρᴬ ρᴮ ρᴬᴮ → rel 𝓣ᴿ σ tᴮ tᴬᴮ → 58 | 𝓔ᴿ (th^Env Synᴬ.th^𝓣 ρᴬ extend ∙ Synᴬ.zro) (ρᴮ ∙ tᴮ) (ρᴬᴮ ∙ tᴬᴮ) 59 | th^𝓔ᴿ : 𝓔ᴿ ρᴬ ρᴮ ρᴬᴮ → (ρ : Thinning Θ Ω) → 60 | 𝓔ᴿ ρᴬ (th^Env Synᴮ.th^𝓣 ρᴮ ρ) (th^Env Synᴬᴮ.th^𝓣 ρᴬᴮ ρ) 61 | \end{code} 62 | % 63 | %<*var> 64 | \begin{code} 65 | varᴿ : 𝓔ᴿ ρᴬ ρᴮ ρᴬᴮ → (v : Var σ Γ) → 𝓡 σ ρᴬ ρᴮ ρᴬᴮ (`var v) 66 | \end{code} 67 | % 68 | %<*zro> 69 | \begin{code} 70 | zroᴿ : rel 𝓣ᴿ σ {σ ∷ Γ} Synᴮ.zro Synᴬᴮ.zro 71 | \end{code} 72 | % 73 | \begin{code} 74 | 75 | private 76 | variable 77 | Synᴬ : Syntactic 𝓣ᴬ 78 | Synᴮ : Syntactic 𝓣ᴮ 79 | Synᴬᴮ : Syntactic 𝓣ᴬᴮ 80 | 𝓔ᴿ : ∀ {Γ Δ Θ} → (Γ ─Env) 𝓣ᴬ Δ → (Δ ─Env) 𝓣ᴮ Θ → (Γ ─Env) 𝓣ᴬᴮ Θ → Set 81 | 𝓣ᴿ : Rel 𝓣ᴮ 𝓣ᴬᴮ 82 | 83 | fromSyn = SynSpec.syntactic 84 | 85 | \end{code} 86 | %<*fundamental> 87 | \begin{code} 88 | module Fundamental (𝓕 : SynFusion Synᴬ Synᴮ Synᴬᴮ 𝓔ᴿ 𝓣ᴿ) where 89 | 90 | open SynFusion 𝓕 91 | 92 | lemma : Fusion (fromSyn Synᴬ) (fromSyn Synᴮ) (fromSyn Synᴬᴮ) 𝓔ᴿ 𝓣ᴿ Eqᴿ 93 | lemma .Fusion.reifyᴬ = id 94 | lemma .Fusion.var0ᴬ = Synᴬ.zro 95 | lemma .Fusion._∙ᴿ_ = _∙ᴿ_ 96 | lemma .Fusion.th^𝓔ᴿ = th^𝓔ᴿ 97 | lemma .Fusion.varᴿ = varᴿ 98 | lemma .Fusion.oneᴿ = λ ρᴿ → refl 99 | lemma .Fusion.ttᴿ = λ ρᴿ → refl 100 | lemma .Fusion.ffᴿ = λ ρᴿ → refl 101 | lemma .Fusion.appᴿ = λ ρᴿ f t → cong₂ `app 102 | lemma .Fusion.ifteᴿ = λ ρᴿ b l r → cong₃ `ifte 103 | lemma .Fusion.lamᴿ = λ ρᴿ b bᴿ → cong `lam (bᴿ extend zroᴿ) 104 | \end{code} 105 | % 106 | -------------------------------------------------------------------------------- /doc/strath.tex: -------------------------------------------------------------------------------- 1 | \documentclass[11pt]{memoir} 2 | 3 | \usepackage{tocbasic} 4 | \DeclareTOCStyleEntry[dynnumwidth]{tocline}{figure}% for figure entries 5 | \DeclareTOCStyleEntry[dynnumwidth]{tocline}{table} 6 | 7 | \usepackage{agda} 8 | \usepackage{catchfilebetweentags} 9 | \usepackage[T1]{fontenc} 10 | \usepackage{mathpartir} 11 | \usepackage{natbib} 12 | 13 | \usepackage{hyperref} 14 | \usepackage{cleveref} 15 | \usepackage{afterpage} 16 | \usepackage{tcolorbox} 17 | %\usepackage{savetrees} 18 | %\usepackage{showframe} 19 | \usepackage{microtype} 20 | 21 | %%%%%%% strathclyde margins 22 | \renewcommand{\baselinestretch}{1.5} 23 | \usepackage{geometry} 24 | \geometry 25 | { lmargin = 4cm 26 | , rmargin = 2.5cm 27 | , tmargin = 2cm 28 | , bmargin = 4cm 29 | } 30 | %%%%%%%%%%%%%%%%%%%%%%%%%%% 31 | 32 | %%%%%%% ligatures 33 | %% disable the >> ligatures 34 | \DisableLigatures[>]{} 35 | %% disable the -- ligatures 36 | \DisableLigatures[-]{} 37 | %%%%%%%%%%%%%%%%%%%%%%%%%%% 38 | 39 | \bibliographystyle{abbrvnat} 40 | 41 | \setlength\mathindent{0.5em} 42 | 43 | \usepackage{thmtools} 44 | \usepackage{thm-restate} 45 | \usepackage{xparse} 46 | \input{convention} 47 | 48 | 49 | \newtheorem{definition}{Definition} 50 | \newtheorem{lemma}{Lemma} 51 | \newtheorem{theorem}{Theorem} 52 | \newtheorem{corollary}{Corollary} 53 | \newtheorem{example}{Example} 54 | 55 | %\newtheorem{convention}{Convention} 56 | %\newtheorem{feature}{Feature} 57 | %\newtheorem{technique}{Technique} 58 | \newtheorem{remark}{Remark} 59 | 60 | \newenvironment{agdasnippet}[1][0] 61 | {\ifx{#1}\undefined 62 | \centering 63 | \else 64 | \fi 65 | }{\ignorespacesafterend% 66 | } 67 | 68 | \input{unicode} 69 | \input{commands} 70 | \input{robust-catch} 71 | 72 | \begin{document} 73 | 74 | \frontmatter 75 | %% \input{cover} 76 | %% \cleardoublepage{} 77 | 78 | \input{titlingpage} 79 | 80 | \mainmatter 81 | \pagestyle{ruled} 82 | \begin{KeepFromToc} 83 | \tableofcontents 84 | \end{KeepFromToc} 85 | 86 | 87 | \input{introduction.tex/intro} 88 | %\input{introduction.tex/type-theory} 89 | \input{introduction.tex/category} 90 | \input{introduction.tex/agda} 91 | 92 | 93 | \part{Type and Scope Preserving Programs} 94 | \label{type-scope-semantics} 95 | \input{type-scope-semantics.tex/syntax.tex} 96 | \input{type-scope-semantics.tex/refactoring.tex} 97 | \input{type-scope-semantics.tex/normalization.tex} 98 | \input{type-scope-semantics.tex/cps-translation.tex} 99 | \input{type-scope-semantics.tex/program-conclusion.tex} 100 | 101 | \part{And Their Proofs} 102 | \label{properties} 103 | \input{type-scope-semantics.tex/simulation.tex} 104 | \input{type-scope-semantics.tex/fusion.tex} 105 | 106 | \input{type-scope-semantics.tex/proof-conclusion.tex} 107 | 108 | \part{A Universe of Well Kinded-and-Scoped Syntaxes with Binding, their Programs and their Proofs} 109 | \label{a-universe} 110 | 111 | \input{generic-syntax.tex/a-universe.tex} 112 | \input{generic-syntax.tex/data.tex} 113 | \input{generic-syntax.tex/syntax.tex} 114 | 115 | \input{generic-syntax.tex/safe-programs.tex} 116 | \input{generic-syntax.tex/compiling.tex} 117 | \input{generic-syntax.tex/other-programs.tex} 118 | 119 | \input{generic-syntax.tex/proofs.tex} 120 | 121 | \input{conclusion} 122 | 123 | \cleardoublepage{} 124 | 125 | 126 | \appendix 127 | 128 | \chapter{Conventions and Techniques}\label{sec:conventions} 129 | \allconventions 130 | \alltechniques 131 | 132 | \chapter{Agda Features}\label{sec:features} 133 | \allfeatures 134 | 135 | \clearpage 136 | \listoffigures 137 | \bibliography{main} 138 | 139 | \end{document} 140 | -------------------------------------------------------------------------------- /doc/convention.tex: -------------------------------------------------------------------------------- 1 | \declaretheorem[ 2 | name=Convention, 3 | ]{thmconvene} 4 | 5 | \ExplSyntaxOn 6 | \NewDocumentEnvironment{convention}{o} 7 | { 8 | \int_gincr:N \g_thesis_convention_int 9 | \IfNoValueTF{#1} 10 | { 11 | \thesis_restatable:x { \int_to_Alph:n { \g_thesis_convention_int } } 12 | } 13 | { 14 | \thesis_restatable:nx { #1 } { \int_to_Alph:n { \g_thesis_convention_int } } 15 | } 16 | } 17 | { 18 | \endrestatable 19 | \seq_gput_right:Nx \g_thesis_convention_seq 20 | { 21 | convene \int_to_Alph:n { \g_thesis_convention_int } 22 | } 23 | } 24 | 25 | \NewDocumentCommand{\allconventions}{} 26 | { 27 | \seq_map_inline:Nn \g_thesis_convention_seq 28 | { 29 | \use:c { ##1 } * 30 | } 31 | } 32 | 33 | \int_new:N \g_thesis_convention_int 34 | \seq_new:N \g_thesis_convention_seq 35 | 36 | \cs_new_protected:Npn \thesis_restatable:n #1 37 | { 38 | \restatable{thmconvene}{convene#1} 39 | } 40 | \cs_generate_variant:Nn \thesis_restatable:n { x } 41 | \cs_new_protected:Npn \thesis_restatable:nn #1 #2 42 | { 43 | \restatable[#1]{thmconvene}{convene#2} 44 | } 45 | \cs_generate_variant:Nn \thesis_restatable:nn { nx } 46 | \ExplSyntaxOff 47 | 48 | \declaretheorem[ 49 | name=Feature, 50 | ]{thmfeaturing} 51 | 52 | \ExplSyntaxOn 53 | \NewDocumentEnvironment{feature}{o} 54 | { 55 | \int_gincr:N \g_thesisb_feature_int 56 | \IfNoValueTF{#1} 57 | { 58 | \thesisb_restatable:x { \int_to_Alph:n { \g_thesisb_feature_int } } 59 | } 60 | { 61 | \thesisb_restatable:nx { #1 } { \int_to_Alph:n { \g_thesisb_feature_int } } 62 | } 63 | } 64 | { 65 | \endrestatable 66 | \seq_gput_right:Nx \g_thesisb_feature_seq 67 | { 68 | featuring \int_to_Alph:n { \g_thesisb_feature_int } 69 | } 70 | } 71 | 72 | \NewDocumentCommand{\allfeatures}{} 73 | { 74 | \seq_map_inline:Nn \g_thesisb_feature_seq 75 | { 76 | \use:c { ##1 } * 77 | } 78 | } 79 | 80 | \int_new:N \g_thesisb_feature_int 81 | \seq_new:N \g_thesisb_feature_seq 82 | 83 | \cs_new_protected:Npn \thesisb_restatable:n #1 84 | { 85 | \restatable{thmfeaturing}{featuring#1} 86 | } 87 | \cs_generate_variant:Nn \thesisb_restatable:n { x } 88 | \cs_new_protected:Npn \thesisb_restatable:nn #1 #2 89 | { 90 | \restatable[#1]{thmfeaturing}{featuring#2} 91 | } 92 | \cs_generate_variant:Nn \thesisb_restatable:nn { nx } 93 | \ExplSyntaxOff 94 | 95 | 96 | \declaretheorem[ 97 | name=Technique, 98 | ]{thmtechnique} 99 | 100 | \ExplSyntaxOn 101 | \NewDocumentEnvironment{technique}{o} 102 | { 103 | \int_gincr:N \g_thesisc_technique_int 104 | \IfNoValueTF{#1} 105 | { 106 | \thesisc_restatable:x { \int_to_Alph:n { \g_thesisc_technique_int } } 107 | } 108 | { 109 | \thesisc_restatable:nx { #1 } { \int_to_Alph:n { \g_thesisc_technique_int } } 110 | } 111 | } 112 | { 113 | \endrestatable 114 | \seq_gput_right:Nx \g_thesisc_technique_seq 115 | { 116 | technique \int_to_Alph:n { \g_thesisc_technique_int } 117 | } 118 | } 119 | 120 | \NewDocumentCommand{\alltechniques}{} 121 | { 122 | \seq_map_inline:Nn \g_thesisc_technique_seq 123 | { 124 | \use:c { ##1 } * 125 | } 126 | } 127 | 128 | \int_new:N \g_thesisc_technique_int 129 | \seq_new:N \g_thesisc_technique_seq 130 | 131 | \cs_new_protected:Npn \thesisc_restatable:n #1 132 | { 133 | \restatable{thmtechnique}{technique#1} 134 | } 135 | \cs_generate_variant:Nn \thesisc_restatable:n { x } 136 | \cs_new_protected:Npn \thesisc_restatable:nn #1 #2 137 | { 138 | \restatable[#1]{thmtechnique}{technique#2} 139 | } 140 | \cs_generate_variant:Nn \thesisc_restatable:nn { nx } 141 | \ExplSyntaxOff 142 | -------------------------------------------------------------------------------- /doc/generic-syntax.agda/Generic/Examples/SystemF.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | {-# OPTIONS --safe --sized-types #-} 3 | 4 | module Generic.Examples.SystemF where 5 | 6 | open import Size 7 | open import Data.Unit 8 | open import Data.Bool 9 | open import Data.Product 10 | open import Data.List.Base hiding ([_]) 11 | open import Function 12 | open import Relation.Binary.PropositionalEquality hiding ([_]) 13 | 14 | open import Data.Var hiding (_<$>_) 15 | open import Data.Var.Varlike 16 | open import Data.Environment hiding (_<$>_) 17 | open import Generic.Syntax 18 | open import Generic.Semantics 19 | open import Generic.Semantics.Syntactic 20 | 21 | data Kind : Set where 22 | Term Type : Kind 23 | 24 | system-F : Desc Kind 25 | system-F = `X [] Type (`X [] Type (`∎ Type)) -- arrow 26 | `+ `X (Type ∷ []) Type (`∎ Type) -- forall 27 | `+ `X [] Term (`X [] Term (`∎ Term)) -- app 28 | `+ `X (Term ∷ []) Term (`∎ Term) -- lam 29 | `+ `X [] Term (`X [] Type (`∎ Term)) -- App 30 | `+ `X (Type ∷ []) Term (`∎ Term) -- Lam 31 | 32 | pattern arr a b = `con (true , a , b , refl) 33 | pattern for b = `con (false , true , b , refl) 34 | 35 | pattern app f t = `con (false , false , true , f , t , refl) 36 | pattern lam b = `con (false , false , false , true , b , refl) 37 | 38 | pattern App f t = `con (false , false , false , false , true , f , t , refl) 39 | pattern Lam b = `con (false , false , false , false , false , b , refl) 40 | 41 | SF : Kind → List Kind → Set 42 | SF = Tm system-F ∞ 43 | 44 | data Redex {Γ : List Kind} : SF Term Γ → Set where 45 | applam : (b : SF Term (Term ∷ Γ)) (u : SF Term Γ) → Redex (app (lam b) u) 46 | AppLam : (b : SF Term (Type ∷ Γ)) (u : SF Type Γ) → Redex (App (Lam b) u) 47 | -- congruence rules 48 | [app-l] : {f : SF Term Γ} → Redex f → (t : SF Term Γ) → Redex (app f t) 49 | [app-r] : (f : SF Term Γ) {t : SF Term Γ} → Redex t → Redex (app f t) 50 | [lam] : {b : SF Term (Term ∷ Γ)} → Redex b → Redex (lam b) 51 | [App] : {f : SF Term Γ} → Redex f → (t : SF Type Γ) → Redex (App f t) 52 | [Lam] : {b : SF Term (Type ∷ Γ)} → Redex b → Redex (Lam b) 53 | 54 | open import Category.Monad 55 | open import Data.Maybe 56 | open import Data.Maybe.Categorical 57 | import Level 58 | open RawMonadPlus (monadPlus {Level.zero}) 59 | 60 | redex : {Γ : List Kind} (t : SF Term Γ) → Maybe (Redex t) 61 | redex (app (lam b) u) = just (applam b u) 62 | redex (App (Lam b) u) = just (AppLam b u) 63 | redex (app f t) = flip [app-l] t <$> redex f ∣ [app-r] f <$> redex t 64 | redex (lam b) = [lam] <$> redex b 65 | redex (App f t) = flip [App] t <$> redex f 66 | redex (Lam b) = [Lam] <$> redex b 67 | redex _ = nothing 68 | 69 | fire : {Γ : List Kind} {t : SF Term Γ} → Redex t → SF Term Γ 70 | fire (applam b u) = sub (base vl^Tm ∙ u) b 71 | fire (AppLam b u) = sub (base vl^Tm ∙ u) b 72 | fire ([app-l] f t) = app (fire f) t 73 | fire ([app-r] f t) = app f (fire t) 74 | fire ([lam] b) = lam (fire b) 75 | fire ([App] f t) = App (fire f) t 76 | fire ([Lam] b) = Lam (fire b) 77 | 78 | open import Codata.Thunk 79 | open import Codata.Colist 80 | open import Codata.Colist.Bisimilarity using (_⊢_≈_; _∷_; []) 81 | 82 | 83 | eval : ∀ {i} (Γ : List Kind) (t : SF Term Γ) → Colist (SF Term Γ) i 84 | eval Γ t with redex t 85 | ... | just r = t ∷ λ where .force → eval Γ (fire r) 86 | ... | nothing = t ∷ λ where .force → [] 87 | 88 | `id : SF Term [] 89 | `id = Lam (lam (`var z)) 90 | 91 | _ : _ ⊢ eval [] (lam (lam (app (lam (app (`var z) (`var z))) (lam (`var z))))) 92 | ≈ fromList (_ ∷ _ ∷ lam (lam (lam (`var z))) ∷ []) 93 | _ = refl ∷ λ where .force → refl ∷ λ where .force → refl ∷ λ where .force → [] 94 | \end{code} 95 | -------------------------------------------------------------------------------- /doc/generic-syntax.agda/Generic/Relator.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | {-# OPTIONS --safe --sized-types #-} 3 | 4 | module Generic.Relator where 5 | 6 | open import Axiom.UniquenessOfIdentityProofs.WithK 7 | open import Size 8 | open import Data.Unit 9 | open import Data.List hiding ([_] ; zip) 10 | open import Data.Product hiding (zip) 11 | open import Relation.Binary.PropositionalEquality hiding ([_]) 12 | 13 | open import Function 14 | open import Relation.Unary 15 | open import Data.Var 16 | open import Data.Var.Varlike 17 | open import Data.Environment 18 | open import Generic.Syntax 19 | open import Generic.Semantics 20 | 21 | private 22 | variable 23 | I : Set 24 | T X Y Z : List I → I ─Scoped 25 | γ δ θ : List I 26 | σ τ : I 27 | 𝓥 𝓦 𝓒 : I ─Scoped 28 | 29 | \end{code} 30 | %<*ziptype> 31 | \begin{code} 32 | ⟦_⟧ᴿ : (d : Desc I) → (∀ Δ σ → ∀[ X Δ σ ⇒ Y Δ σ ⇒ const Set ]) 33 | → ∀[ ⟦ d ⟧ X σ ⇒ ⟦ d ⟧ Y σ ⇒ const Set ] 34 | ⟦ `∎ j ⟧ᴿ R x y = ⊤ 35 | ⟦ `X Δ j d ⟧ᴿ R (r , x) (r' , y) = R Δ j r r' × ⟦ d ⟧ᴿ R x y 36 | ⟦ `σ A d ⟧ᴿ R (a , x) (a' , y) = Σ (a' ≡ a) (λ where refl → ⟦ d a ⟧ᴿ R x y) 37 | \end{code} 38 | % 39 | \begin{code} 40 | 41 | module _ 42 | {R : ∀ θ σ → ∀[ X θ σ ⇒ Y θ σ ⇒ const Set ]} 43 | {f : ∀ θ σ → T θ σ γ → X θ σ δ} 44 | {g : ∀ θ σ → T θ σ γ → Y θ σ δ} 45 | where 46 | 47 | liftᴿ : ∀ d (FG : ∀ θ σ t → R θ σ (f θ σ t) (g θ σ t)) → 48 | ∀ (t : ⟦ d ⟧ T σ γ) → ⟦ d ⟧ᴿ R (fmap d f t) (fmap d g t) 49 | liftᴿ (`σ A d) FG (a , t) = refl , liftᴿ (d a) FG t 50 | liftᴿ (`X j Δ d) FG (x , t) = FG j Δ x , liftᴿ d FG t 51 | liftᴿ (`∎ j) FG refl = tt 52 | 53 | module _ {R : ∀ θ σ → ∀[ X θ σ ⇒ X θ σ ⇒ const Set ]} where 54 | 55 | reflᴿ : ∀ d (re : ∀ θ σ (x : X θ σ γ) → R θ σ x x) → 56 | (t : ⟦ d ⟧ X σ γ) → ⟦ d ⟧ᴿ R t t 57 | reflᴿ (`σ A d) re (a , t) = refl , reflᴿ (d a) re t 58 | reflᴿ (`X j Δ d) re (x , t) = re j Δ x , reflᴿ d re t 59 | reflᴿ (`∎ j) re refl = tt 60 | 61 | module _ {R : ∀ θ σ → ∀[ X θ σ ⇒ Y θ σ ⇒ const Set ]} 62 | {S : ∀ θ σ → ∀[ Y θ σ ⇒ X θ σ ⇒ const Set ]} 63 | where 64 | 65 | symᴿ : ∀ d (sy : ∀ θ σ {x : X θ σ γ} {y} → R θ σ x y → S θ σ y x) → 66 | ∀ {t : ⟦ d ⟧ X σ γ} {u} → ⟦ d ⟧ᴿ R t u → ⟦ d ⟧ᴿ S u t 67 | symᴿ (`σ A d) sy (refl , t) = refl , symᴿ (d _) sy t 68 | symᴿ (`X j Δ d) sy (r , t) = sy j Δ r , symᴿ d sy t 69 | symᴿ (`∎ j) sy tt = tt 70 | 71 | module _ {R : ∀ θ σ → ∀[ X θ σ ⇒ Y θ σ ⇒ const Set ]} 72 | {S : ∀ θ σ → ∀[ Y θ σ ⇒ Z θ σ ⇒ const Set ]} 73 | {RS : ∀ θ σ → ∀[ X θ σ ⇒ Z θ σ ⇒ const Set ]} 74 | where 75 | 76 | transᴿ : ∀ d (tr : ∀ θ σ {x : X θ σ γ} {y z} → R θ σ x y → S θ σ y z → RS θ σ x z) → 77 | ∀ {t : ⟦ d ⟧ X σ γ} {u v} → ⟦ d ⟧ᴿ R t u → ⟦ d ⟧ᴿ S u v → ⟦ d ⟧ᴿ RS t v 78 | transᴿ (`σ A d) tr (refl , t) (refl , u) = refl , transᴿ (d _) tr t u 79 | transᴿ (`X j Δ d) tr (x , t) (y , u) = tr j Δ x y , transᴿ d tr t u 80 | transᴿ (`∎ j) tr tt tt = tt 81 | 82 | 83 | open import Data.Relation 84 | 85 | module _ (𝓥ᴿ : Rel {I} 𝓥 𝓦) {vl^𝓥 : VarLike 𝓥} {vl^𝓦 : VarLike 𝓦} where 86 | 87 | reifyᴿ : ∀ d (re : ∀ θ σ {v : Kripke 𝓥 𝓒 θ σ γ} {w} → 88 | Kripkeᴿ 𝓥ᴿ Eqᴿ θ σ v w → 89 | reify vl^𝓥 θ σ v ≡ reify vl^𝓦 θ σ w) → 90 | ∀ {bv : ⟦ d ⟧ (Kripke 𝓥 𝓒) σ γ} {bw : ⟦ d ⟧ (Kripke 𝓦 𝓒) σ γ} → 91 | ⟦ d ⟧ᴿ (Kripkeᴿ 𝓥ᴿ Eqᴿ) bv bw → 92 | (⟦ d ⟧ (Scope 𝓒) σ γ ∋ fmap d (reify vl^𝓥) bv) ≡ fmap d (reify vl^𝓦) bw 93 | reifyᴿ (`σ A d) re (refl , t) = cong -,_ (reifyᴿ (d _) re t) 94 | reifyᴿ (`X j Δ d) re (x , t) = cong₂ _,_ (re j Δ x) (reifyᴿ d re t) 95 | reifyᴿ (`∎ j) re tt = uip _ _ 96 | \end{code} 97 | -------------------------------------------------------------------------------- /doc/type-scope-semantics.agda/Semantics/NormalisationByEvaluation/BetaIotaXi.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | module Semantics.NormalisationByEvaluation.BetaIotaXi where 3 | 4 | open import Data.Unit using (⊤) 5 | open import Data.Bool using (Bool; true; false; if_then_else_) 6 | open import Data.Sum as Sum using (_⊎_; inj₁; inj₂) 7 | open import Data.List.Base using (List; []; _∷_) 8 | open import Data.Var 9 | open import Data.Environment hiding (Kripke) 10 | open import Syntax.Type 11 | open import Syntax.Calculus 12 | open import Relation.Unary 13 | open import Relation.Binary.PropositionalEquality using (_≡_; refl) 14 | open import Function 15 | 16 | \end{code} 17 | %<*noeta> 18 | \begin{code} 19 | NoEta^βιξ : Type → Set 20 | NoEta^βιξ _ = ⊤ 21 | \end{code} 22 | % 23 | \begin{code} 24 | 25 | open import Syntax.Normal NoEta^βιξ 26 | open import Syntax.Normal.Thinnable 27 | open import Semantics.Specification 28 | open import Semantics.NormalisationByEvaluation.Specification 29 | 30 | private 31 | 32 | variable 33 | 34 | σ τ : Type 35 | Γ Δ : List Type 36 | 37 | \end{code} 38 | %<*model> 39 | \begin{code} 40 | mutual 41 | 42 | Model : Type ─Scoped 43 | Model σ Γ = Ne σ Γ ⊎ Value σ Γ 44 | 45 | Value : Type ─Scoped 46 | Value `Unit = const ⊤ 47 | Value `Bool = const Bool 48 | Value (σ `→ τ) = □ (Model σ ⇒ Model τ) 49 | \end{code} 50 | % 51 | %<*thmodel> 52 | \begin{code} 53 | th^Value : ∀ σ → Thinnable (Value σ) 54 | th^Value `Unit = th^const 55 | th^Value `Bool = th^const 56 | th^Value (σ `→ τ) = th^□ 57 | 58 | th^Model : ∀ σ → Thinnable (Model σ) 59 | th^Model σ (inj₁ neu) ρ = inj₁ (th^Ne neu ρ) 60 | th^Model σ (inj₂ val) ρ = inj₂ (th^Value σ val ρ) 61 | \end{code} 62 | % 63 | \begin{code} 64 | 65 | \end{code} 66 | %<*reflect> 67 | \begin{code} 68 | reflect : ∀[ Ne σ ⇒ Model σ ] 69 | reflect = inj₁ 70 | 71 | var0 : ∀[ (σ ∷_) ⊢ Model σ ] 72 | var0 = reflect (`var z) 73 | \end{code} 74 | % 75 | %<*reify> 76 | \begin{code} 77 | mutual 78 | 79 | reify : ∀ σ → ∀[ Model σ ⇒ Nf σ ] 80 | reify σ (inj₁ neu) = `neu _ neu 81 | reify σ (inj₂ val) = reify^Value σ val 82 | 83 | reify^Value : ∀ σ → ∀[ Value σ ⇒ Nf σ ] 84 | reify^Value `Unit _ = `one 85 | reify^Value `Bool b = if b then `tt else `ff 86 | reify^Value (σ `→ τ) f = `lam (reify τ (f extend var0)) 87 | \end{code} 88 | % 89 | \begin{code} 90 | module _ {σ τ} where 91 | \end{code} 92 | %<*app> 93 | \begin{code} 94 | APP : ∀[ Model (σ `→ τ) ⇒ Model σ ⇒ Model τ ] 95 | APP (inj₁ f) t = inj₁ (`app f (reify σ t)) 96 | APP (inj₂ f) t = extract f t 97 | \end{code} 98 | % 99 | \begin{code} 100 | module _ {σ} where 101 | \end{code} 102 | %<*ifte> 103 | \begin{code} 104 | IFTE : ∀[ Model `Bool ⇒ Model σ ⇒ Model σ ⇒ Model σ ] 105 | IFTE (inj₁ b) l r = inj₁ (`ifte b (reify σ l) (reify σ r)) 106 | IFTE (inj₂ b) l r = if b then l else r 107 | \end{code} 108 | % 109 | \begin{code} 110 | open Semantics 111 | 112 | \end{code} 113 | %<*eval> 114 | \begin{code} 115 | Eval : Semantics Model Model 116 | Eval .th^𝓥 = th^Model _ 117 | Eval .var = id 118 | Eval .lam = inj₂ 119 | Eval .app = APP 120 | Eval .one = inj₂ _ 121 | Eval .tt = inj₂ true 122 | Eval .ff = inj₂ false 123 | Eval .ifte = IFTE 124 | \end{code} 125 | % 126 | %<*norm> 127 | \begin{code} 128 | nbe : NBE Model Nf 129 | nbe = record 130 | { Sem = Eval 131 | ; embed = reflect ∘ `var 132 | ; reify = reify _ 133 | } 134 | \end{code} 135 | % 136 | 137 | \begin{code} 138 | open NBE using (test) 139 | \end{code} 140 | 141 | %<*test> 142 | \begin{code} 143 | _ : test nbe ≡ `lam (`lam (`neu _ (`ifte (`var (s z)) `one (`neu _ (`var z))))) 144 | _ = refl 145 | \end{code} 146 | % 147 | -------------------------------------------------------------------------------- /doc/type-scope-semantics.agda/Syntax/Normal.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | open import Syntax.Type 3 | 4 | module Syntax.Normal (NoEta : Type → Set) where 5 | 6 | open import Data.Environment 7 | open import Data.Relation 8 | open import Syntax.Calculus 9 | open import Data.List.Base using (List; []; _∷_) 10 | open import Syntax.Type 11 | open import Data.Var 12 | open import Function 13 | open import Relation.Unary 14 | open import Relation.Binary.PropositionalEquality 15 | open import Relation.Binary.PropositionalEquality.Extra 16 | 17 | private 18 | 19 | variable 20 | 21 | σ τ : Type 22 | Γ Δ Θ : List Type 23 | 24 | 25 | \end{code} 26 | %<*normal> 27 | \begin{code} 28 | mutual 29 | 30 | data Ne : Type ─Scoped where 31 | `var : ∀[ Var σ ⇒ Ne σ ] 32 | `app : ∀[ Ne (σ `→ τ) ⇒ Nf σ ⇒ Ne τ ] 33 | `ifte : ∀[ Ne `Bool ⇒ Nf σ ⇒ Nf σ ⇒ Ne σ ] 34 | 35 | data Nf : Type ─Scoped where 36 | `neu : NoEta σ → ∀[ Ne σ ⇒ Nf σ ] 37 | `one : ∀[ Nf `Unit ] 38 | `tt `ff : ∀[ Nf `Bool ] 39 | `lam : ∀[ (σ ∷_) ⊢ Nf τ ⇒ Nf (σ `→ τ) ] 40 | \end{code} 41 | % 42 | \begin{code} 43 | `neu-injective : ∀ {p q} {t u : Ne σ Γ} → `neu p t ≡ `neu q u → t ≡ u 44 | `neu-injective refl = refl 45 | 46 | erase^Ne : ∀[ Ne σ ⇒ Term σ ] 47 | erase^Nf : ∀[ Nf σ ⇒ Term σ ] 48 | 49 | erase^Ne (`var v) = `var v 50 | erase^Ne (`app f t) = `app (erase^Ne f) (erase^Nf t) 51 | erase^Ne (`ifte b l r) = `ifte (erase^Ne b) (erase^Nf l) (erase^Nf r) 52 | 53 | erase^Nf (`neu _ t) = erase^Ne t 54 | erase^Nf `one = `one 55 | erase^Nf `tt = `tt 56 | erase^Nf `ff = `ff 57 | erase^Nf (`lam b) = `lam (erase^Nf b) 58 | 59 | th^Ne : Thinnable (Ne σ) 60 | th^Nf : Thinnable (Nf σ) 61 | 62 | th^Ne (`var v) ρ = `var (th^Var v ρ) 63 | th^Ne (`app f t) ρ = `app (th^Ne f ρ) (th^Nf t ρ) 64 | th^Ne (`ifte b l r) ρ = `ifte (th^Ne b ρ) (th^Nf l ρ) (th^Nf r ρ) 65 | 66 | th^Nf (`neu ne t) ρ = `neu ne (th^Ne t ρ) 67 | th^Nf `one ρ = `one 68 | th^Nf `tt ρ = `tt 69 | th^Nf `ff ρ = `ff 70 | th^Nf (`lam b) ρ = `lam (th^Nf b (select ρ extend ∙ z)) 71 | 72 | th^Ne-id : ∀ (t : Ne σ Γ) {ρ} → All Eqᴿ Γ ρ (pack id) → th^Ne t ρ ≡ t 73 | th^Nf-id : ∀ (t : Nf σ Γ) {ρ} → All Eqᴿ Γ ρ (pack id) → th^Nf t ρ ≡ t 74 | 75 | th^Ne-id (`var v) ρᴿ = cong `var (lookupᴿ ρᴿ v) 76 | th^Ne-id (`app f t) ρᴿ = cong₂ `app (th^Ne-id f ρᴿ) (th^Nf-id t ρᴿ) 77 | th^Ne-id (`ifte b l r) ρᴿ = cong₃ `ifte (th^Ne-id b ρᴿ) (th^Nf-id l ρᴿ) (th^Nf-id r ρᴿ) 78 | 79 | th^Nf-id (`neu ne t) ρᴿ = cong (`neu ne) (th^Ne-id t ρᴿ) 80 | th^Nf-id `one ρᴿ = refl 81 | th^Nf-id `tt ρᴿ = refl 82 | th^Nf-id `ff ρᴿ = refl 83 | th^Nf-id (`lam b) ρᴿ = cong `lam $ th^Nf-id b $ packᴿ λ where 84 | z → refl 85 | (s v) → cong s (lookupᴿ ρᴿ v) 86 | 87 | th^Ne-compose : 88 | (t : Ne σ Γ) {ρ₁ : Thinning Γ Δ} {ρ₂ : Thinning Δ Θ} {ρ₃ : Thinning Γ Θ} → 89 | All Eqᴿ Γ (select ρ₁ ρ₂) ρ₃ → th^Ne (th^Ne t ρ₁) ρ₂ ≡ th^Ne t ρ₃ 90 | th^Nf-compose : 91 | (t : Nf σ Γ) {ρ₁ : Thinning Γ Δ} {ρ₂ : Thinning Δ Θ} {ρ₃ : Thinning Γ Θ} → 92 | All Eqᴿ Γ (select ρ₁ ρ₂) ρ₃ → th^Nf (th^Nf t ρ₁) ρ₂ ≡ th^Nf t ρ₃ 93 | 94 | th^Ne-compose (`var v) ρᴿ = cong `var (lookupᴿ ρᴿ v) 95 | th^Ne-compose (`app f t) ρᴿ = cong₂ `app (th^Ne-compose f ρᴿ) (th^Nf-compose t ρᴿ) 96 | th^Ne-compose (`ifte b l r) ρᴿ = 97 | cong₃ `ifte (th^Ne-compose b ρᴿ) (th^Nf-compose l ρᴿ) (th^Nf-compose r ρᴿ) 98 | 99 | th^Nf-compose (`neu ne t) ρᴿ = cong (`neu ne) (th^Ne-compose t ρᴿ) 100 | th^Nf-compose `one ρᴿ = refl 101 | th^Nf-compose `tt ρᴿ = refl 102 | th^Nf-compose `ff ρᴿ = refl 103 | th^Nf-compose (`lam b) ρᴿ = cong `lam $ th^Nf-compose b $ packᴿ λ where 104 | z → refl 105 | (s v) → cong s (lookupᴿ ρᴿ v) 106 | 107 | \end{code} 108 | -------------------------------------------------------------------------------- /doc/type-scope-semantics.agda/Semantics/NormalisationByEvaluation/BetaIota.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | module Semantics.NormalisationByEvaluation.BetaIota where 3 | 4 | open import Data.Unit using (⊤) 5 | open import Data.Bool using (Bool; true; false; if_then_else_) 6 | open import Data.Sum as Sum using (_⊎_; inj₁; inj₂) 7 | open import Data.List.Base using (List; []; _∷_) 8 | open import Data.Product 9 | 10 | open import Data.Var 11 | open import Data.Environment hiding (Kripke) 12 | open import Syntax.Type 13 | open import Syntax.Calculus 14 | open import Semantics.Syntactic.Instances 15 | open import Relation.Unary 16 | open import Function 17 | 18 | open import Syntax.WeakHead 19 | open import Semantics.Specification 20 | open import Semantics.NormalisationByEvaluation.Specification 21 | 22 | open import Agda.Builtin.Equality 23 | 24 | private 25 | 26 | variable 27 | 28 | σ τ : Type 29 | 30 | \end{code} 31 | %<*model> 32 | \begin{code} 33 | mutual 34 | 35 | Model : Type ─Scoped 36 | Model σ Γ = Term σ Γ × (WHNE σ Γ ⊎ Value σ Γ) 37 | 38 | Value : Type ─Scoped 39 | Value `Unit = const ⊤ 40 | Value `Bool = const Bool 41 | Value (σ `→ τ) = □ (Model σ ⇒ Model τ) 42 | \end{code} 43 | % 44 | %<*thmodel> 45 | \begin{code} 46 | th^Value : ∀ σ → Thinnable (Value σ) 47 | th^Value `Unit = th^const 48 | th^Value `Bool = th^const 49 | th^Value (σ `→ τ) = th^□ 50 | 51 | th^Model : ∀ σ → Thinnable (Model σ) 52 | th^Model σ (t , inj₁ whne) ρ = th^Term t ρ , inj₁ (th^WHNE whne ρ) 53 | th^Model σ (t , inj₂ val) ρ = th^Term t ρ , inj₂ (th^Value σ val ρ) 54 | \end{code} 55 | % 56 | \begin{code} 57 | 58 | \end{code} 59 | %<*reifyreflect> 60 | \begin{code} 61 | reflect : ∀[ WHNE σ ⇒ Model σ ] 62 | reflect = < erase^WHNE , inj₁ > 63 | 64 | var0 : ∀[ (σ ∷_) ⊢ Model σ ] 65 | var0 = reflect (`var z) 66 | 67 | mutual 68 | 69 | reify : ∀[ Model σ ⇒ WHNF σ ] 70 | reify (t , inj₁ whne) = `whne whne 71 | reify (t , inj₂ val) = reify^Value _ val 72 | 73 | reify^Value : ∀ σ → ∀[ Value σ ⇒ WHNF σ ] 74 | reify^Value `Unit v = `one 75 | reify^Value `Bool v = if v then `tt else `ff 76 | reify^Value (σ `→ τ) v = `lam (proj₁ (v extend var0)) 77 | \end{code} 78 | % 79 | %<*app> 80 | \begin{code} 81 | APP : ∀[ Model (σ `→ τ) ⇒ Model σ ⇒ Model τ ] 82 | APP (f , inj₁ whne) (t , _) = (`app f t , inj₁ (`app whne t)) 83 | APP (f , inj₂ fun) T@(t , _) = (`app f t , proj₂ (extract fun T)) 84 | \end{code} 85 | % 86 | %<*ifte> 87 | \begin{code} 88 | IFTE : ∀[ Model `Bool ⇒ Model σ ⇒ Model σ ⇒ Model σ ] 89 | IFTE (b , inj₁ whne) (l , _) (r , _) = (`ifte b l r , inj₁ (`ifte whne l r)) 90 | IFTE (b , inj₂ v) l r = if v then l else r 91 | \end{code} 92 | % 93 | %<*lam> 94 | \begin{code} 95 | LAM : ∀[ □ (Model σ ⇒ Model τ) ⇒ Model (σ `→ τ) ] 96 | LAM b = (`lam (proj₁ (b extend var0)) , inj₂ b) 97 | \end{code} 98 | % 99 | \begin{code} 100 | open Semantics 101 | 102 | \end{code} 103 | %<*eval> 104 | \begin{code} 105 | Eval : Semantics Model Model 106 | Eval .th^𝓥 = th^Model _ 107 | Eval .var = id 108 | Eval .lam = LAM 109 | Eval .app = APP 110 | Eval .one = `one , inj₂ _ 111 | Eval .tt = `tt , inj₂ true 112 | Eval .ff = `ff , inj₂ false 113 | Eval .ifte = IFTE 114 | \end{code} 115 | % 116 | %<*whnorm> 117 | \begin{code} 118 | nbe : NBE Model WHNF 119 | nbe = record 120 | { Sem = Eval 121 | ; embed = reflect ∘ `var 122 | ; reify = reify 123 | } 124 | \end{code} 125 | % 126 | 127 | 128 | \begin{code} 129 | open NBE using (test) 130 | \end{code} 131 | 132 | %<*test> 133 | \begin{code} 134 | _ : test nbe ≡ `lam (`lam (`app (`lam (`var z)) 135 | (`ifte (`var (s z)) `one (`var z)))) 136 | _ = refl 137 | \end{code} 138 | % 139 | 140 | -------------------------------------------------------------------------------- /doc/generic-syntax.agda/Generic/Syntax/LetCounter.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | 3 | {-# OPTIONS --safe --sized-types #-} 4 | 5 | module Generic.Syntax.LetCounter where 6 | 7 | open import Algebra 8 | open import Algebra.Structures 9 | open import Data.Bool 10 | open import Data.List.Relation.Unary.All as All 11 | open import Data.Product using (_×_; _,_) 12 | open import Agda.Builtin.List 13 | open import Agda.Builtin.Equality 14 | open import Relation.Unary 15 | open import Function 16 | 17 | open import Data.Var 18 | open import Generic.Syntax 19 | 20 | open import Generic.Syntax.LetBinder using (Let) 21 | 22 | module _ {a} {A : Set a} (expensive : A) where 23 | 24 | -- inlining is hard 25 | 26 | _ : A × A 27 | -- here we better not inline x (but it's fine to inline y) 28 | \end{code} 29 | %<*expensive> 30 | \begin{code} 31 | _ = let x = expensive in 32 | let y = (x , x) in 33 | y 34 | \end{code} 35 | % 36 | \begin{code} 37 | _ : A 38 | -- here on the other hand we can inline all the lets! 39 | \end{code} 40 | %<*cheap> 41 | \begin{code} 42 | _ = let x = expensive in 43 | let y = (x , x) in 44 | x 45 | \end{code} 46 | % 47 | 48 | %<*counter> 49 | \begin{code} 50 | data Counter : Set where 51 | zero : Counter 52 | one : Counter 53 | many : Counter 54 | \end{code} 55 | % 56 | %<*addition> 57 | \begin{code} 58 | _+_ : Counter → Counter → Counter 59 | zero + n = n 60 | m + zero = m 61 | _ + _ = many 62 | \end{code} 63 | % 64 | 65 | %<*multiplication> 66 | \begin{code} 67 | _*_ : Counter → Counter → Counter 68 | zero * n = zero 69 | m * zero = zero 70 | one * n = n 71 | m * one = m 72 | many * many = many 73 | \end{code} 74 | % 75 | 76 | \begin{code} 77 | 78 | module _ {I : Set} where 79 | 80 | private 81 | variable 82 | σ : I 83 | 84 | \end{code} 85 | %<*count> 86 | \begin{code} 87 | Count : List I → Set 88 | Count = All (const Counter) 89 | \end{code} 90 | % 91 | %<*zeros> 92 | \begin{code} 93 | zeros : ∀[ Count ] 94 | zeros {[]} = [] 95 | zeros {σ ∷ Γ} = zero ∷ zeros 96 | \end{code} 97 | % 98 | %<*fromVar> 99 | \begin{code} 100 | fromVar : ∀[ Var σ ⇒ Count ] 101 | fromVar z = one ∷ zeros 102 | fromVar (s v) = zero ∷ fromVar v 103 | \end{code} 104 | % 105 | \begin{code} 106 | 107 | \end{code} 108 | %<*merge> 109 | \begin{code} 110 | merge : ∀[ Count ⇒ Count ⇒ Count ] 111 | merge [] [] = [] 112 | merge (m ∷ cs) (n ∷ ds) = 113 | (m + n) ∷ merge cs ds 114 | \end{code} 115 | % 116 | \begin{code} 117 | 118 | \end{code} 119 | %<*control> 120 | \begin{code} 121 | control : Counter → ∀[ Count ⇒ Count ] 122 | control zero cs = zeros 123 | control one cs = cs -- inlined 124 | control many cs = cs -- not inlined 125 | \end{code} 126 | % 127 | \begin{code} 128 | 129 | \end{code} 130 | %<*scale> 131 | \begin{code} 132 | scale : Counter → ∀[ Count ⇒ Count ] 133 | scale zero cs = zeros 134 | scale one cs = cs 135 | scale k cs = map (k *_) cs 136 | \end{code} 137 | % 138 | \begin{code} 139 | 140 | 141 | rawMonoid : List I → RawMonoid _ _ 142 | rawMonoid Γ = record 143 | { Carrier = Count Γ 144 | ; _≈_ = _≡_ 145 | ; _∙_ = merge 146 | ; ε = tabulate (λ _ → zero) 147 | } 148 | 149 | \end{code} 150 | %<*clet> 151 | \begin{code} 152 | CLet : Desc I 153 | CLet = `σ Counter $ λ _ → Let 154 | \end{code} 155 | % 156 | \begin{code} 157 | 158 | 159 | pattern `IN' e t = (_ , e , t , refl) 160 | pattern `IN e t = `con (`IN' e t) 161 | 162 | module _ {I : Set} {d : Desc I} where 163 | 164 | embed : ∀ {i σ} → ∀[ Tm d i σ ⇒ Tm (d `+ CLet) i σ ] 165 | embed = map^Tm (MkDescMorphism (true ,_)) 166 | 167 | \end{code} 168 | -------------------------------------------------------------------------------- /doc/type-scope-semantics.agda/Properties/Simulation/Specification.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | open import Syntax.Type 3 | open import Data.Var 4 | 5 | module Properties.Simulation.Specification {𝓥ᴬ 𝓒ᴬ 𝓥ᴮ 𝓒ᴮ : Type ─Scoped} where 6 | 7 | open import Data.Environment 8 | open import Data.List.Base using (List; []; _∷_) 9 | open import Data.Relation 10 | open import Syntax.Type 11 | open import Syntax.Calculus 12 | open import Semantics.Specification 13 | open import Function renaming (_$′_ to _$_) using () 14 | 15 | private 16 | variable 17 | σ τ : Type 18 | Γ Δ Θ : List Type 19 | ρᴬ : (Γ ─Env) 𝓥ᴬ Δ 20 | ρᴮ : (Γ ─Env) 𝓥ᴮ Δ 21 | vᴬ : 𝓥ᴬ σ Γ 22 | vᴮ : 𝓥ᴮ σ Γ 23 | 24 | \end{code} 25 | %<*simulation> 26 | \begin{code} 27 | record Simulation (𝓢ᴬ : Semantics 𝓥ᴬ 𝓒ᴬ) (𝓢ᴮ : Semantics 𝓥ᴮ 𝓒ᴮ) 28 | (𝓥ᴿ : Rel 𝓥ᴬ 𝓥ᴮ) (𝓒ᴿ : Rel 𝓒ᴬ 𝓒ᴮ) : Set where 29 | \end{code} 30 | % 31 | \begin{code} 32 | module 𝓢ᴬ = Semantics 𝓢ᴬ 33 | module 𝓢ᴮ = Semantics 𝓢ᴮ 34 | \end{code} 35 | %<*crel> 36 | \begin{code} 37 | 𝓡 : ∀ σ → (Γ ─Env) 𝓥ᴬ Δ → (Γ ─Env) 𝓥ᴮ Δ → Term σ Γ → Set 38 | 𝓡 σ ρᴬ ρᴮ t = rel 𝓒ᴿ σ (semantics 𝓢ᴬ ρᴬ t) (semantics 𝓢ᴮ ρᴮ t) 39 | \end{code} 40 | % 41 | %<*rkripke> 42 | \begin{code} 43 | Kripkeᴿ : ∀ {Γ Δ} σ τ → (Γ ─Env) 𝓥ᴬ Δ → (Γ ─Env) 𝓥ᴮ Δ → 44 | Term τ (σ ∷ Γ) → Set 45 | Kripkeᴿ {Γ} {Δ} σ τ ρᴬ ρᴮ b = 46 | ∀ {Θ} (ρ : Thinning Δ Θ) {uᴬ uᴮ} → 𝓥ᴿ .rel σ uᴬ uᴮ → 47 | 𝓡 τ (th^Env 𝓢ᴬ.th^𝓥 ρᴬ ρ ∙ uᴬ) (th^Env 𝓢ᴮ.th^𝓥 ρᴮ ρ ∙ uᴮ) b 48 | \end{code} 49 | % 50 | \begin{code} 51 | field 52 | \end{code} 53 | %<*thV> 54 | \begin{code} 55 | th^𝓥ᴿ : (ρ : Thinning Δ Θ) → 𝓥ᴿ .rel σ vᴬ vᴮ → 𝓥ᴿ .rel σ (𝓢ᴬ.th^𝓥 vᴬ ρ) (𝓢ᴮ.th^𝓥 vᴮ ρ) 56 | \end{code} 57 | % 58 | %<*var> 59 | \begin{code} 60 | varᴿ : All 𝓥ᴿ Γ ρᴬ ρᴮ → (v : Var σ Γ) → 𝓡 σ ρᴬ ρᴮ (`var v) 61 | \end{code} 62 | % 63 | %<*lam> 64 | \begin{code} 65 | lamᴿ : All 𝓥ᴿ Γ ρᴬ ρᴮ → ∀ b → Kripkeᴿ σ τ ρᴬ ρᴮ b → 𝓡 (σ `→ τ) ρᴬ ρᴮ (`lam b) 66 | \end{code} 67 | % 68 | %<*struct> 69 | \begin{code} 70 | appᴿ : All 𝓥ᴿ Γ ρᴬ ρᴮ → 71 | ∀ f t → 𝓡 (σ `→ τ) ρᴬ ρᴮ f → 𝓡 σ ρᴬ ρᴮ t → 72 | 𝓡 τ ρᴬ ρᴮ (`app f t) 73 | ifteᴿ : All 𝓥ᴿ Γ ρᴬ ρᴮ → 74 | ∀ b l r → 𝓡 `Bool ρᴬ ρᴮ b → 𝓡 σ ρᴬ ρᴮ l → 𝓡 σ ρᴬ ρᴮ r → 75 | 𝓡 σ ρᴬ ρᴮ (`ifte b l r) 76 | \end{code} 77 | % 78 | %<*base> 79 | \begin{code} 80 | oneᴿ : All 𝓥ᴿ Γ ρᴬ ρᴮ → 𝓡 `Unit ρᴬ ρᴮ `one 81 | ttᴿ : All 𝓥ᴿ Γ ρᴬ ρᴮ → 𝓡 `Bool ρᴬ ρᴮ `tt 82 | ffᴿ : All 𝓥ᴿ Γ ρᴬ ρᴮ → 𝓡 `Bool ρᴬ ρᴮ `ff 83 | \end{code} 84 | % 85 | \begin{code} 86 | 87 | 88 | private 89 | variable 90 | 𝓢ᴬ : Semantics 𝓥ᴬ 𝓒ᴬ 91 | 𝓢ᴮ : Semantics 𝓥ᴮ 𝓒ᴮ 92 | 𝓥ᴿ : Rel 𝓥ᴬ 𝓥ᴮ 93 | 𝓒ᴿ : Rel 𝓒ᴬ 𝓒ᴮ 94 | 95 | module _ (𝓢ᴿ : Simulation 𝓢ᴬ 𝓢ᴮ 𝓥ᴿ 𝓒ᴿ) where 96 | 97 | open Simulation 𝓢ᴿ 98 | \end{code} 99 | %<*fundamental:type> 100 | \begin{code} 101 | simulation : All 𝓥ᴿ Γ ρᴬ ρᴮ → ∀ t → 𝓡 σ ρᴬ ρᴮ t 102 | \end{code} 103 | % 104 | %<*fundamental:var> 105 | \begin{code} 106 | simulation ρᴿ (`var v) = varᴿ ρᴿ v 107 | \end{code} 108 | % 109 | %<*fundamental:lam> 110 | \begin{code} 111 | simulation ρᴿ (`lam b) = lamᴿ ρᴿ b $ λ ren vᴿ → 112 | simulation ((th^𝓥ᴿ ren <$>ᴿ ρᴿ) ∙ᴿ vᴿ) b 113 | \end{code} 114 | % 115 | %<*fundamental:base> 116 | \begin{code} 117 | simulation ρᴿ `one = oneᴿ ρᴿ 118 | simulation ρᴿ `tt = ttᴿ ρᴿ 119 | simulation ρᴿ `ff = ffᴿ ρᴿ 120 | \end{code} 121 | % 122 | %<*fundamental:struct> 123 | \begin{code} 124 | simulation ρᴿ (`app f t) = appᴿ ρᴿ f t (simulation ρᴿ f) (simulation ρᴿ t) 125 | simulation ρᴿ (`ifte b l r) = 126 | ifteᴿ ρᴿ b l r (simulation ρᴿ b) (simulation ρᴿ l) (simulation ρᴿ r) 127 | \end{code} 128 | % 129 | -------------------------------------------------------------------------------- /doc/type-scope-semantics.agda/Semantics/NormalisationByEvaluation/BetaIotaXiEta.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | 3 | module Semantics.NormalisationByEvaluation.BetaIotaXiEta where 4 | 5 | open import Data.Unit using (⊤) 6 | open import Data.List.Base using (List; []; _∷_) 7 | open import Data.Var 8 | open import Data.Environment as Env hiding (Kripke; Thinning) 9 | open import Syntax.Type 10 | open import Syntax.Calculus hiding (module DISPLAYONLY) 11 | open import Relation.Unary 12 | open import Function 13 | 14 | \end{code} 15 | %<*noeta> 16 | \begin{code} 17 | data NoEta^βιξη : Type → Set where 18 | `Bool : NoEta^βιξη `Bool 19 | \end{code} 20 | % 21 | \begin{code} 22 | 23 | 24 | open import Syntax.Normal NoEta^βιξη public 25 | open import Syntax.Normal.Thinnable 26 | open import Semantics.Specification 27 | open import Semantics.NormalisationByEvaluation.Specification 28 | 29 | open import Agda.Builtin.Equality 30 | 31 | private 32 | variable 33 | σ τ : Type 34 | Γ : List Type 35 | 36 | module DISPLAYONLY where 37 | \end{code} 38 | %<*model> 39 | \begin{code} 40 | Model : Type ─Scoped 41 | Model `Unit = const ⊤ 42 | Model `Bool = Nf `Bool 43 | Model (σ `→ τ) = □ (Model σ ⇒ Model τ) 44 | \end{code} 45 | % 46 | \begin{code} 47 | Model : Type ─Scoped 48 | Model `Unit Γ = ⊤ 49 | Model `Bool Γ = Nf `Bool Γ 50 | Model (σ `→ τ) Γ = □ (Model σ ⇒ Model τ) Γ 51 | \end{code} 52 | %<*thmodel> 53 | \begin{code} 54 | th^Model : ∀ σ → Thinnable (Model σ) 55 | th^Model `Unit = th^const 56 | th^Model `Bool = th^Nf 57 | th^Model (σ `→ τ) = th^□ 58 | \end{code} 59 | % 60 | %<*reifyreflect> 61 | \begin{code} 62 | mutual 63 | 64 | var0 : ∀[ (σ ∷_) ⊢ Model σ ] 65 | var0 = reflect _ (`var z) 66 | 67 | reflect : ∀ σ → ∀[ Ne σ ⇒ Model σ ] 68 | reflect `Unit t = _ 69 | reflect `Bool t = `neu `Bool t 70 | reflect (σ `→ τ) t = λ ρ u → reflect τ (`app (th^Ne t ρ) (reify σ u)) 71 | 72 | reify : ∀ σ → ∀[ Model σ ⇒ Nf σ ] 73 | reify `Unit T = `one 74 | reify `Bool T = T 75 | reify (σ `→ τ) T = `lam (reify τ (T extend var0)) 76 | \end{code} 77 | % 78 | %<*app> 79 | \begin{code} 80 | APP : ∀[ Model (σ `→ τ) ⇒ Model σ ⇒ Model τ ] 81 | APP t u = extract t u 82 | \end{code} 83 | % 84 | \begin{code} 85 | module _ {σ} where 86 | \end{code} 87 | %<*ifte> 88 | \begin{code} 89 | IFTE : Model `Bool Γ → Model σ Γ → Model σ Γ → Model σ Γ 90 | IFTE `tt l r = l 91 | IFTE `ff l r = r 92 | IFTE (`neu _ T) l r = reflect σ (`ifte T (reify σ l) (reify σ r)) 93 | \end{code} 94 | % 95 | \begin{code} 96 | 97 | open Semantics 98 | \end{code} 99 | %<*eval> 100 | \begin{code} 101 | Eval : Semantics Model Model 102 | Eval .th^𝓥 = th^Model _ 103 | Eval .var = id 104 | Eval .lam = id 105 | Eval .app = APP 106 | Eval .one = _ 107 | Eval .tt = `tt 108 | Eval .ff = `ff 109 | Eval .ifte = IFTE 110 | \end{code} 111 | % 112 | %<*norm> 113 | \begin{code} 114 | nbe : NBE Model Nf 115 | nbe = record 116 | { Sem = Eval 117 | ; embed = reflect _ ∘ `var 118 | ; reify = reify _ 119 | } 120 | \end{code} 121 | % 122 | 123 | \begin{code} 124 | open NBE using (test) 125 | \end{code} 126 | 127 | %<*test> 128 | \begin{code} 129 | _ : test nbe ≡ `lam (`lam `one) 130 | _ = refl 131 | \end{code} 132 | % 133 | %<*exotic> 134 | \begin{code} 135 | exotic : Model (`Bool `→ `Bool) [] 136 | exotic {_ ∷ `Bool ∷ Δ} ρ b = `neu `Bool (`var (s z)) 137 | exotic {_} ρ b = b 138 | \end{code} 139 | % 140 | 141 | \begin{code} 142 | open import Relation.Binary.PropositionalEquality 143 | private 144 | Thinning = Env.Thinning {I = Type} 145 | 146 | 2⇒2 : Type 147 | 2⇒2 = `Bool `→ `Bool 148 | \end{code} 149 | %<*exotictest> 150 | \begin{code} 151 | _ : th^Nf (reify 2⇒2 exotic) (bind `Bool) ≡ `lam (`neu `Bool (`var z)) 152 | _ = refl 153 | 154 | _ : reify 2⇒2 (th^Model 2⇒2 exotic (bind `Bool)) ≡ `lam (`neu `Bool (`var (s z))) 155 | _ = refl 156 | \end{code} 157 | % 158 | -------------------------------------------------------------------------------- /doc/generic-syntax.agda/Generic/Examples/TypeChecking.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | module Generic.Examples.TypeChecking where 3 | 4 | open import Size 5 | open import Function 6 | open import Data.Empty 7 | open import Data.Unit 8 | open import Data.Bool 9 | open import Data.Product as P 10 | open import Data.List hiding ([_]) 11 | open import Data.Maybe as Maybe using (Maybe; nothing; just) 12 | open import Data.Maybe.Categorical as MC 13 | 14 | open import Relation.Unary 15 | open import Data.Var hiding (_<$>_) 16 | open import Data.Environment hiding (_<$>_) 17 | open import Generic.Syntax 18 | open import Generic.Semantics 19 | 20 | import Category.Monad as CM 21 | import Level 22 | module M = CM.RawMonad (MC.monad {Level.zero}) 23 | open M 24 | 25 | open import Relation.Binary.PropositionalEquality hiding ([_]) 26 | 27 | infixr 5 _`→_ 28 | \end{code} 29 | %<*type> 30 | \begin{code} 31 | data Type : Set where 32 | α : Type 33 | _`→_ : Type → Type → Type 34 | \end{code} 35 | % 36 | \begin{code} 37 | 38 | infix 3 _==_ 39 | \end{code} 40 | %<*equal> 41 | \begin{code} 42 | _==_ : (σ τ : Type) → Maybe ⊤ 43 | α == α = just tt 44 | (σ₁ `→ τ₁) == (σ₂ `→ τ₂) = 45 | σ₁ == σ₂ >> τ₁ == τ₂ 46 | _ == _ = nothing 47 | \end{code} 48 | % 49 | %<*arrow> 50 | \begin{code} 51 | isArrow : Type → Maybe (Type × Type) 52 | isArrow (σ `→ τ) = just (σ , τ) 53 | isArrow α = nothing 54 | \end{code} 55 | % 56 | \begin{code} 57 | 58 | 59 | \end{code} 60 | %<*constructors> 61 | \begin{code} 62 | data LangC : Set where 63 | App Lam Emb : LangC 64 | Cut : Type → LangC 65 | \end{code} 66 | % 67 | %<*mode> 68 | \begin{code} 69 | data Mode : Set where 70 | Check Infer : Mode 71 | \end{code} 72 | % 73 | %<*bidirectional> 74 | \begin{code} 75 | Lang : Desc Mode 76 | Lang = `σ LangC $ λ where 77 | App → `X [] Infer (`X [] Check (`∎ Infer)) 78 | Lam → `X (Infer ∷ []) Check (`∎ Check) 79 | (Cut σ) → `X [] Check (`∎ Infer) 80 | Emb → `X [] Infer (`∎ Check) 81 | \end{code} 82 | % 83 | \begin{code} 84 | private 85 | variable 86 | Γ Δ : List Mode 87 | 88 | pattern `app f t = `con (App , f , t , refl) 89 | pattern `lam b = `con (Lam , b , refl) 90 | pattern `cut σ t = `con (Cut σ , t , refl) 91 | pattern `emb t = `con (Emb , t , refl) 92 | 93 | \end{code} 94 | %<*typemode> 95 | \begin{code} 96 | Type- : Mode → Set 97 | Type- Check = Type → Maybe ⊤ 98 | Type- Infer = Maybe Type 99 | \end{code} 100 | % 101 | %<*varmode> 102 | \begin{code} 103 | data Var- : Mode → Set where 104 | `var : Type → Var- Infer 105 | \end{code} 106 | % 107 | \begin{code} 108 | open Semantics 109 | 110 | \end{code} 111 | %<*app> 112 | \begin{code} 113 | APP : Type- Infer → Type- Check → Type- Infer 114 | APP f t = do 115 | σ`→τ ← f 116 | (σ , τ) ← isArrow σ`→τ 117 | τ <$ t σ 118 | \end{code} 119 | % 120 | %<*lam> 121 | \begin{code} 122 | LAM : Kripke (const ∘ Var-) (const ∘ Type-) (Infer ∷ []) Check Γ → Type- Check 123 | LAM b σ`→τ = do 124 | (σ , τ) ← isArrow σ`→τ 125 | b (bind Infer) (ε ∙ `var σ) τ 126 | \end{code} 127 | % 128 | %<*typecheck> 129 | \begin{code} 130 | Typecheck : Semantics Lang (const ∘ Var-) (const ∘ Type-) 131 | Typecheck .th^𝓥 = th^const 132 | Typecheck .var = λ where (`var t) → just t 133 | Typecheck .alg = λ where 134 | (App , f , t , refl) → APP f t 135 | (Lam , b , refl) → LAM b 136 | (Cut σ , t , refl) → σ <$ t σ 137 | (Emb , t , refl) → λ σ → t >>= σ ==_ 138 | \end{code} 139 | % 140 | \begin{code} 141 | private variable m : Mode 142 | 143 | \end{code} 144 | %<*typefun> 145 | \begin{code} 146 | type- : ∀ m → TM Lang m → Type- m 147 | type- m t = Semantics.closed Typecheck t 148 | \end{code} 149 | % 150 | %<*example> 151 | \begin{code} 152 | _ : let `id : TM Lang Check 153 | `id = `lam (`emb (`var z)) 154 | in type- Infer (`app (`cut ((α `→ α) `→ (α `→ α)) `id) `id) ≡ just (α `→ α) 155 | _ = refl 156 | \end{code} 157 | % 158 | \begin{code} 159 | -------------------------------------------------------------------------------- /doc/generic-syntax.agda/Generic/Simulation.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | {-# OPTIONS --safe --sized-types #-} 3 | 4 | open import Data.Var hiding (_<$>_; z; s) 5 | open import Data.Relation 6 | 7 | module Generic.Simulation {I : Set} {𝓥ᴬ 𝓥ᴮ 𝓒ᴬ 𝓒ᴮ : I ─Scoped} where 8 | 9 | open import Size 10 | open import Data.List hiding ([_] ; lookup ; zip) 11 | open import Function 12 | open import Relation.Binary.PropositionalEquality hiding ([_]) 13 | 14 | open import Relation.Unary 15 | 16 | open import Data.Var.Varlike 17 | open import Data.Environment 18 | open import Generic.Syntax 19 | open import Generic.Semantics 20 | open import Generic.Relator as Relator using (⟦_⟧ᴿ; liftᴿ) 21 | 22 | private 23 | variable 24 | Γ Δ : List I 25 | σ : I 26 | vᴬ : 𝓥ᴬ σ Γ 27 | vᴮ : 𝓥ᴮ σ Γ 28 | s : Size 29 | ρᴬ : (Γ ─Env) 𝓥ᴬ Δ 30 | ρᴮ : (Γ ─Env) 𝓥ᴮ Δ 31 | 32 | module _ (𝓥ᴿ : Rel 𝓥ᴬ 𝓥ᴮ) (𝓒ᴿ : Rel 𝓒ᴬ 𝓒ᴮ) where 33 | 34 | reifyᴿ : {vlᴬ : VarLike 𝓥ᴬ} {vlᴮ : VarLike 𝓥ᴮ} (vlᴿ : VarLikeᴿ 𝓥ᴿ vlᴬ vlᴮ) → 35 | ∀ Δ σ → {kᴬ : Kripke 𝓥ᴬ 𝓒ᴬ Δ σ Γ} {kᴮ : Kripke 𝓥ᴮ 𝓒ᴮ Δ σ Γ} → 36 | Kripkeᴿ 𝓥ᴿ 𝓒ᴿ Δ σ kᴬ kᴮ → rel 𝓒ᴿ σ (reify vlᴬ Δ σ kᴬ) (reify vlᴮ Δ σ kᴮ) 37 | reifyᴿ vlᴿ [] σ kᴿ = kᴿ 38 | reifyᴿ vlᴿ Δ@(_ ∷ _) σ kᴿ = kᴿ (freshʳ vl^Var Δ) (VarLikeᴿ.freshˡᴿ vlᴿ _) 39 | 40 | private 41 | module DISPLAYONLY where 42 | \end{code} 43 | %<*recsim1> 44 | \begin{code} 45 | record Simulation (d : Desc I) 46 | (𝓢ᴬ : Semantics d 𝓥ᴬ 𝓒ᴬ) (𝓢ᴮ : Semantics d 𝓥ᴮ 𝓒ᴮ) 47 | (𝓥ᴿ : Rel 𝓥ᴬ 𝓥ᴮ) (𝓒ᴿ : Rel 𝓒ᴬ 𝓒ᴮ) : Set where 48 | module 𝓢ᴬ = Semantics 𝓢ᴬ 49 | module 𝓢ᴮ = Semantics 𝓢ᴮ 50 | field thᴿ : (ρ : Thinning Γ Δ) → rel 𝓥ᴿ σ vᴬ vᴮ → 51 | rel 𝓥ᴿ σ (𝓢ᴬ.th^𝓥 vᴬ ρ) (𝓢ᴮ.th^𝓥 vᴮ ρ) 52 | varᴿ : rel 𝓥ᴿ σ vᴬ vᴮ → rel 𝓒ᴿ σ (𝓢ᴬ.var vᴬ) (𝓢ᴮ.var vᴮ) 53 | bodyᴿ : ⟦ d ⟧ (Kripke 𝓥ᴬ 𝓒ᴬ) σ Δ → ⟦ d ⟧ (Kripke 𝓥ᴮ 𝓒ᴮ) σ Δ → Set 54 | bodyᴿ vᴬ vᴮ = ⟦ d ⟧ᴿ (Kripkeᴿ 𝓥ᴿ 𝓒ᴿ) vᴬ vᴮ 55 | field algᴿ : (b : ⟦ d ⟧ (Scope (Tm d s)) σ Γ) → All 𝓥ᴿ Γ ρᴬ ρᴮ → 56 | let vᴬ = fmap d (𝓢ᴬ.body ρᴬ) b 57 | vᴮ = fmap d (𝓢ᴮ.body ρᴮ) b 58 | in bodyᴿ vᴬ vᴮ → rel 𝓒ᴿ σ (𝓢ᴬ.alg vᴬ) (𝓢ᴮ.alg vᴮ) 59 | \end{code} 60 | % 61 | 62 | %<*recsim> 63 | \begin{code} 64 | record Simulation (d : Desc I) 65 | (𝓢ᴬ : Semantics d 𝓥ᴬ 𝓒ᴬ) (𝓢ᴮ : Semantics d 𝓥ᴮ 𝓒ᴮ) 66 | (𝓥ᴿ : Rel 𝓥ᴬ 𝓥ᴮ) (𝓒ᴿ : Rel 𝓒ᴬ 𝓒ᴮ) : Set where 67 | \end{code} 68 | % 69 | \begin{code} 70 | module 𝓢ᴬ = Semantics 𝓢ᴬ 71 | module 𝓢ᴮ = Semantics 𝓢ᴮ 72 | field 73 | \end{code} 74 | %<*thR> 75 | \begin{code} 76 | thᴿ : (ρ : Thinning Γ Δ) → rel 𝓥ᴿ σ vᴬ vᴮ → rel 𝓥ᴿ σ (𝓢ᴬ.th^𝓥 vᴬ ρ) (𝓢ᴮ.th^𝓥 vᴮ ρ) 77 | \end{code} 78 | % 79 | %<*varR> 80 | \begin{code} 81 | varᴿ : rel 𝓥ᴿ σ vᴬ vᴮ → rel 𝓒ᴿ σ (𝓢ᴬ.var vᴬ) (𝓢ᴮ.var vᴮ) 82 | \end{code} 83 | % 84 | %<*bodyR> 85 | \begin{code} 86 | bodyᴿ : ⟦ d ⟧ (Kripke 𝓥ᴬ 𝓒ᴬ) σ Δ → ⟦ d ⟧ (Kripke 𝓥ᴮ 𝓒ᴮ) σ Δ → Set 87 | bodyᴿ vᴬ vᴮ = ⟦ d ⟧ᴿ (Kripkeᴿ 𝓥ᴿ 𝓒ᴿ) vᴬ vᴮ 88 | \end{code} 89 | % 90 | \begin{code} 91 | field 92 | \end{code} 93 | %<*algR> 94 | \begin{code} 95 | algᴿ : (b : ⟦ d ⟧ (Scope (Tm d s)) σ Γ) → All 𝓥ᴿ Γ ρᴬ ρᴮ → 96 | let vᴬ = fmap d (𝓢ᴬ.body ρᴬ) b 97 | vᴮ = fmap d (𝓢ᴮ.body ρᴮ) b 98 | in bodyᴿ vᴬ vᴮ → rel 𝓒ᴿ σ (𝓢ᴬ.alg vᴬ) (𝓢ᴮ.alg vᴮ) 99 | \end{code} 100 | % 101 | %<*simbody> 102 | %<*simtype> 103 | \begin{code} 104 | sim : All 𝓥ᴿ Γ ρᴬ ρᴮ → (t : Tm d s σ Γ) → 105 | rel 𝓒ᴿ σ (𝓢ᴬ.semantics ρᴬ t) (𝓢ᴮ.semantics ρᴮ t) 106 | \end{code} 107 | % 108 | \begin{code} 109 | body : All 𝓥ᴿ Γ ρᴬ ρᴮ → ∀ Δ j → (t : Scope (Tm d s) Δ j Γ) → 110 | Kripkeᴿ 𝓥ᴿ 𝓒ᴿ Δ j (𝓢ᴬ.body ρᴬ Δ j t) (𝓢ᴮ.body ρᴮ Δ j t) 111 | 112 | sim ρᴿ (`var k) = varᴿ (lookupᴿ ρᴿ k) 113 | sim ρᴿ (`con t) = algᴿ t ρᴿ (liftᴿ d (body ρᴿ) t) 114 | 115 | body ρᴿ [] i t = sim ρᴿ t 116 | body ρᴿ (_ ∷ _) i t = λ σ vsᴿ → sim (vsᴿ ++^Envᴿ (thᴿ σ <$>ᴿ ρᴿ)) t 117 | \end{code} 118 | % 119 | -------------------------------------------------------------------------------- /doc/generic-syntax.tex/proofs.tex: -------------------------------------------------------------------------------- 1 | \chapter{Building Generic Proofs about Generic Programs} 2 | 3 | We have already shown in \cref{sec:simulationrel,sec:fusionrel} that, for the 4 | simply typed $\lambda$-calculus, introducing an abstract notion of Semantics 5 | not only reveals the shared structure of common traversals, it also allows 6 | us to give abstract proof frameworks for simulation or fusion lemmas. These 7 | ideas naturally extend to our generic presentation of semantics for all syntaxes. 8 | 9 | \section{Additional Relation Transformers} 10 | 11 | During our exploration of generic proofs about the behaviour of \AR{Semantics} 12 | for a concrete object language, we have introduced a notion \AR{Rel} of 13 | relations as well as a relation transformer for environments 14 | (cf. \cref{sec:relation-transformers}). Working on a universe of syntaxes, 15 | we are going to need to define additional relators. 16 | 17 | \paragraph{Kripke relator} 18 | The Kripke relator is a generalisation of the ad-hoc definition introduced 19 | in \cref{fig:relationalkripke}. We assume that we have two types of values 20 | \AB{𝓥ᴬ} and \AB{𝓥ᴮ} 21 | as well as a relation \AB{𝓥ᴿ} for pairs of such values, and two types 22 | of computations \AB{𝓒ᴬ} and \AB{𝓒ᴮ} whose notion of relatedness is 23 | given by \AB{𝓒ᴿ}. We can define 24 | \AF{Kripkeᴿ} relating Kripke functions of type 25 | {(\AF{Kripke} \AB{𝓥ᴬ} \AB{𝓒ᴬ})} and {(\AF{Kripke} \AB{𝓥ᴮ} \AB{𝓒ᴮ})} 26 | respectively by stating that they send related inputs 27 | to related outputs. In this definition we use the relation transformer for 28 | environment called \AF{All} and introduced in \cref{defn:Env-rel}. 29 | 30 | \begin{figure}[h] 31 | \ExecuteMetaData[shared/Data/Var/Varlike.tex]{kripkeR} 32 | \caption{Relational Kripke Function Spaces: From Related Inputs to Related Outputs\label{fig:Kripke-rel}} 33 | \end{figure} 34 | 35 | \paragraph{Desc relator}\label{def:zipd} 36 | The relator (\AF{⟦} \AB{d} \AF{⟧ᴿ}) is a relation transformer which characterises 37 | structurally equal layers such that their substructures are themselves related 38 | by the relation it is passed as an argument. It inherits a lot of its relational 39 | arguments' properties: whenever \AB{R} is reflexive (respectively symmetric or 40 | transitive) so is {(\AF{⟦} \AB{d} \AF{⟧ᴿ} \AB{R})}.\label{lem:zipstable} 41 | 42 | It is defined by induction on the description and case analysis on the two 43 | layers which are meant to be equal: 44 | \begin{itemize} 45 | \item In the stop token case {(\AIC{`∎} \AB{i})}, the two layers are two proofs 46 | that the branches' respective indices match the specified one. We consider 47 | these two proofs to be trivially equal (i.e. the constraint generated is the 48 | unit type). 49 | % 50 | This would not hold true in Homotopy Type Theory~(\cite{hottbook}) 51 | but if we were to generalise the work to that setting, we could either explicitly 52 | restrict our setup to language whose indices are equipped with a decidable equality 53 | or insist on studying the ways in which these proofs can be equal. 54 | \item When facing a recursive position {(\AIC{`X} \AB{Δ} \AB{j} \AB{d})}, we 55 | demand that the two substructures are related by {\AB{R} \AB{Δ} \AB{j}} 56 | and that the rest of the layers are related by (\AF{⟦} \AB{d} \AF{⟧ᴿ} \AB{R}) 57 | \item Two nodes of type {(\AIC{`σ} \AB{A} \AB{d})} will 58 | be related if they both carry the same payload \AB{a} of type \AB{A} and if 59 | the rest of the layers are related by (\AF{⟦} \AB{d} \AB{a} \AF{⟧ᴿ} \AB{R}) 60 | \end{itemize} 61 | 62 | \begin{figure}[h] 63 | \ExecuteMetaData[generic-syntax.agda/Generic/Relator.tex]{ziptype} 64 | \caption{Relator: Characterising Structurally Equal Values with Related Substructures\label{fig:zip-rel}} 65 | \end{figure} 66 | 67 | If we were to take a fixpoint of \AF{⟦\_⟧ᴿ}, we could obtain a structural 68 | notion of equality for terms which we could prove equivalent to propositional 69 | equality. Although interesting in its own right, this section will focus 70 | on more advanced use-cases. 71 | 72 | %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% 73 | %% SIMULATION 74 | 75 | \input{generic-syntax.tex/proofs/simulation.tex} 76 | \input{generic-syntax.tex/proofs/fusion.tex} 77 | -------------------------------------------------------------------------------- /doc/generic-syntax.agda/Generic/Semantics/Printing.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | {-# OPTIONS --safe --sized-types #-} 3 | 4 | module Generic.Semantics.Printing {I : Set} where 5 | 6 | open import Codata.Stream using (Stream) 7 | open import Size 8 | open import Data.Product 9 | open import Data.List.Base using (List; []; _∷_; _++_) 10 | open import Data.String using (String) 11 | open import Category.Monad 12 | open import Category.Monad.State 13 | open import Function 14 | open import Relation.Unary 15 | 16 | -- We reuse Name, Printer, M, fresh, and names from the STLC printing example 17 | 18 | open import Semantics.Printing 19 | using (Fresh; Wrap; Name; Printer; MkW; getW; map^Wrap; th^Wrap; fresh; names) 20 | 21 | private 22 | variable 23 | Γ Δ : List I 24 | σ : I 25 | i : Size 26 | 27 | -- The Printing Monad we are working with: a state containing a stream 28 | -- of *distinct* Strings. 29 | module ST = RawMonadState (StateMonadState (Stream String _)) 30 | open ST renaming (rawIApplicative to ApplicativeM) 31 | hiding (_<$>_) 32 | 33 | open import Data.Var hiding (get; _<$>_) 34 | open import Data.Environment hiding (_>>_; sequenceA; _<$>_) 35 | open import Data.Var.Varlike 36 | open import Generic.Syntax hiding (sequenceA) 37 | open import Generic.Semantics 38 | 39 | \end{code} 40 | %<*vlmname> 41 | \begin{code} 42 | vl^FreshName : VarLike (λ (σ : I) → Fresh ∘ (Name σ)) 43 | vl^FreshName = record 44 | { th^𝓥 = th^Functor functor^Fresh th^Wrap 45 | ; new = fresh _ 46 | } 47 | \end{code} 48 | % 49 | \begin{code} 50 | where open ST renaming (rawFunctor to functor^Fresh) 51 | 52 | -- To print a term the user need to explain to us how to display one 53 | -- layer of term given that the newly-bound variables have been assigned 54 | -- fresh names and the subterms have already been rendered using these 55 | -- names. 56 | 57 | \end{code} 58 | %<*pieces> 59 | \begin{code} 60 | Pieces : List I → I ─Scoped 61 | Pieces [] i Γ = String 62 | Pieces Δ i Γ = (Δ ─Env) Name (Δ ++ Γ) × String 63 | \end{code} 64 | % 65 | %<*reifytype> 66 | \begin{code} 67 | reify^Pieces : ∀ Δ i → Kripke Name Printer Δ i Γ → Fresh (Pieces Δ i Γ) 68 | \end{code} 69 | % 70 | %<*reifybase> 71 | \begin{code} 72 | reify^Pieces [] i p = getW p 73 | \end{code} 74 | % 75 | %<*reifypieces> 76 | \begin{code} 77 | reify^Pieces Δ@(_ ∷ _) i f = do 78 | ρ ← sequenceA (freshˡ vl^FreshName _) 79 | b ← getW (f (freshʳ vl^Var Δ) ρ) 80 | return (ρ , b) 81 | \end{code} 82 | % 83 | \begin{code} 84 | where open Data.Environment 85 | instance _ = ApplicativeM 86 | 87 | \end{code} 88 | %<*display> 89 | \begin{code} 90 | Display : Desc I → Set 91 | Display d = ∀ {i Γ} → ⟦ d ⟧ Pieces i Γ → String 92 | \end{code} 93 | % 94 | \begin{code} 95 | 96 | --------------------------------------------------------------------- 97 | -- Generic Printing Semantics 98 | 99 | -- Given a strategy to `Display` one layer of term we can generate a full 100 | -- printer. 101 | 102 | open Semantics 103 | 104 | module _ {d : Desc I} where 105 | 106 | \end{code} 107 | %<*printing> 108 | \begin{code} 109 | Printing : Display d → Semantics d Name Printer 110 | Printing dis .th^𝓥 = th^Wrap 111 | Printing dis .var = map^Wrap return 112 | Printing dis .alg = λ v → MkW $ dis <$> mapA d reify^Pieces v 113 | \end{code} 114 | % 115 | \begin{code} 116 | where open Generic.Syntax 117 | open ST 118 | instance _ = ApplicativeM 119 | 120 | -- Corollary: a generic printer using a silly name supply 121 | 122 | 123 | open Data.Environment 124 | instance _ = ApplicativeM 125 | 126 | \end{code} 127 | %<*print> 128 | \begin{AgdaAlign} 129 | \AgdaNoSpaceAroundCode 130 | %<*printtype> 131 | \begin{code} 132 | print : Display d → Tm d i σ Γ → String 133 | \end{code} 134 | % 135 | \begin{AgdaSuppressSpace} 136 | \begin{code} 137 | print dis t = proj₁ (printer names) where 138 | printer : Fresh String 139 | printer = do 140 | init ← sequenceA (base vl^FreshName) 141 | getW (Semantics.semantics (Printing dis) init t) 142 | \end{code} 143 | \AgdaSpaceAroundCode 144 | \end{AgdaSuppressSpace} 145 | \end{AgdaAlign} 146 | % 147 | \begin{code} 148 | -------------------------------------------------------------------------------- /doc/generic-syntax.agda/Motivation/Problem/Naive.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | {-# OPTIONS --safe --sized-types #-} 3 | 4 | module Motivation.Problem.Naive where 5 | 6 | open import Relation.Unary 7 | open import Data.Var hiding (_<$>_) 8 | open import Data.Environment 9 | open import Data.Relation 10 | 11 | open import Function 12 | open import Relation.Binary.PropositionalEquality hiding ([_]) 13 | open ≡-Reasoning 14 | open import Data.List hiding ([_] ; lookup) 15 | open import Data.Star 16 | 17 | open import Syntax.Type 18 | 19 | private 20 | variable 21 | σ τ : Type 22 | Γ Δ : List Type 23 | 24 | \end{code} 25 | %<*source> 26 | \begin{code} 27 | data S : Type ─Scoped where 28 | `var : ∀[ Var σ ⇒ S σ ] 29 | `lam : ∀[ (σ ∷_) ⊢ S τ ⇒ S (σ `→ τ) ] 30 | `app : ∀[ S (σ `→ τ) ⇒ S σ ⇒ S τ ] 31 | `let : ∀[ S σ ⇒ (σ ∷_) ⊢ S τ ⇒ S τ ] 32 | \end{code} 33 | % 34 | %<*target> 35 | \begin{code} 36 | data T : Type ─Scoped where 37 | `var : ∀[ Var σ ⇒ T σ ] 38 | `lam : ∀[ (σ ∷_) ⊢ T τ ⇒ T (σ `→ τ) ] 39 | `app : ∀[ T (σ `→ τ) ⇒ T σ ⇒ T τ ] 40 | \end{code} 41 | % 42 | %<*thT> 43 | \begin{code} 44 | th^T : Thinnable (T σ) 45 | th^T (`var v) ρ = `var (lookup ρ v) 46 | th^T (`lam b) ρ = `lam (th^T b (th^Env th^Var ρ extend ∙ z)) 47 | th^T (`app f t) ρ = `app (th^T f ρ) (th^T t ρ) 48 | \end{code} 49 | % 50 | %<*unlet> 51 | \begin{code} 52 | unlet : (Γ ─Env) T Δ → S σ Γ → T σ Δ 53 | unlet ρ (`var v) = lookup ρ v 54 | unlet ρ (`lam b) = `lam (unlet (th^Env th^T ρ extend ∙ `var z) b) 55 | unlet ρ (`app f t) = `app (unlet ρ f) (unlet ρ t) 56 | unlet ρ (`let e t) = unlet (ρ ∙ unlet ρ e) t 57 | \end{code} 58 | % 59 | %<*thS> 60 | \begin{code} 61 | th^S : Thinnable (S σ) 62 | th^S (`var v) ρ = `var (lookup ρ v) 63 | th^S (`lam b) ρ = `lam (th^S b (th^Env th^Var ρ extend ∙ z)) 64 | th^S (`app f t) ρ = `app (th^S f ρ) (th^S t ρ) 65 | th^S (`let e t) ρ = `let (th^S e ρ) (th^S t (th^Env th^Var ρ extend ∙ z)) 66 | \end{code} 67 | % 68 | %<*subS> 69 | \begin{code} 70 | sub^S : (Γ ─Env) S Δ → S σ Γ → S σ Δ 71 | sub^S ρ (`var v) = lookup ρ v 72 | sub^S ρ (`lam b) = `lam (sub^S (th^Env th^S ρ extend ∙ `var z) b) 73 | sub^S ρ (`app f t) = `app (sub^S ρ f) (sub^S ρ t) 74 | sub^S ρ (`let e t) = `let (sub^S ρ e) (sub^S (th^Env th^S ρ extend ∙ `var z) t) 75 | \end{code} 76 | % 77 | %<*subT> 78 | \begin{code} 79 | sub^T : (Γ ─Env) T Δ → T σ Γ → T σ Δ 80 | sub^T ρ (`var v) = lookup ρ v 81 | sub^T ρ (`lam b) = `lam (sub^T (th^Env th^T ρ extend ∙ `var z) b) 82 | sub^T ρ (`app f t) = `app (sub^T ρ f) (sub^T ρ t) 83 | 84 | _⟨_/0⟩^S : ∀[ (σ ∷_) ⊢ S τ ⇒ S σ ⇒ S τ ] 85 | b ⟨ t /0⟩^S = sub^S (`var <$> identity ∙ t) b 86 | 87 | _⟨_/0⟩^T : ∀[ (σ ∷_) ⊢ T τ ⇒ T σ ⇒ T τ ] 88 | b ⟨ t /0⟩^T = sub^T (`var <$> identity ∙ t) b 89 | \end{code} 90 | % 91 | \begin{code} 92 | 93 | module _ where 94 | infix 1 _⊢_∋_↝S_ 95 | 96 | private 97 | variable 98 | b c : S τ (σ ∷ Γ) 99 | f g : S (σ `→ τ) Γ 100 | d e t u : S σ Γ 101 | \end{code} 102 | %<*redS> 103 | \begin{code} 104 | data _⊢_∋_↝S_ : ∀ Γ σ → S σ Γ → S σ Γ → Set where 105 | -- computation 106 | β : ∀ (b : S τ (σ ∷ Γ)) u → Γ ⊢ τ ∋ `app (`lam b) u ↝S b ⟨ u /0⟩^S 107 | ζ : ∀ e (t : S τ (σ ∷ Γ)) → Γ ⊢ τ ∋ `let e t ↝S t ⟨ e /0⟩^S 108 | -- structural 109 | `lam : (σ ∷ Γ) ⊢ τ ∋ b ↝S c → Γ ⊢ σ `→ τ ∋ `lam b ↝S `lam c 110 | `appl : Γ ⊢ σ `→ τ ∋ f ↝S g → ∀ t → Γ ⊢ τ ∋ `app f t ↝S `app g t 111 | `appr : ∀ f → Γ ⊢ σ ∋ t ↝S u → Γ ⊢ τ ∋ `app f t ↝S `app f u 112 | `letl : Γ ⊢ σ ∋ d ↝S e → ∀ t → Γ ⊢ τ ∋ `let d t ↝S `let e t 113 | `letr : ∀ e → (σ ∷ Γ) ⊢ τ ∋ t ↝S u → Γ ⊢ τ ∋ `let e t ↝S `let e u 114 | \end{code} 115 | % 116 | \begin{code} 117 | module _ where 118 | infix 1 _⊢_∋_↝T_ 119 | 120 | private 121 | variable 122 | b c : T τ (σ ∷ Γ) 123 | f g : T (σ `→ τ) Γ 124 | d e t u : T σ Γ 125 | \end{code} 126 | \begin{code} 127 | data _⊢_∋_↝T_ : ∀ Γ σ → T σ Γ → T σ Γ → Set where 128 | -- computation 129 | β : ∀ {Γ σ τ} (b : T τ (σ ∷ Γ)) u → Γ ⊢ τ ∋ `app (`lam b) u ↝T b ⟨ u /0⟩^T 130 | -- structural 131 | `lam : (σ ∷ Γ) ⊢ τ ∋ b ↝T c → Γ ⊢ σ `→ τ ∋ `lam b ↝T `lam c 132 | `appl : Γ ⊢ σ `→ τ ∋ f ↝T g → ∀ t → Γ ⊢ τ ∋ `app f t ↝T `app g t 133 | `appr : ∀ f → Γ ⊢ σ ∋ t ↝T u → Γ ⊢ τ ∋ `app f t ↝T `app f u 134 | \end{code} 135 | -------------------------------------------------------------------------------- /doc/generic-syntax.agda/Generic/Fusion/Specialised/Propositional.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | -------------------------------------------------------------------------------- 3 | -- This is the closest one can come to an axiom-free verison of Kaiser, Schäfer, 4 | -- and Stark's observation that our notion of Semantics is intrinsically able to 5 | -- absord any Renaming which may have come before. 6 | -- 7 | -- This is used both to replicate Kaiser, Schäfer, and Stark's result in the 8 | -- module Generic.Fusion.Specialised and to make the fusion proofs of renaming 9 | -- with renaming, substitution, and let-elaboration simpler. 10 | -------------------------------------------------------------------------------- 11 | 12 | {-# OPTIONS --safe --sized-types #-} 13 | 14 | module Generic.Fusion.Specialised.Propositional where 15 | 16 | open import Relation.Unary 17 | open import Data.Var hiding (_<$>_) 18 | open import Data.Environment 19 | open import Data.Var.Varlike 20 | open import Data.Relation 21 | open import Generic.Syntax 22 | open import Generic.Semantics 23 | open import Generic.Semantics.Syntactic 24 | open import Generic.Relator 25 | open import Generic.Fusion 26 | open import Generic.Identity 27 | 28 | open import Size 29 | open import Function 30 | open import Data.Sum 31 | open import Data.Product 32 | open import Data.List.Base hiding (lookup) 33 | open import Relation.Binary.PropositionalEquality hiding ([_]) 34 | open ≡-Reasoning 35 | 36 | module _ {I} (d : Desc I) {𝓥 𝓒} (S : Semantics d 𝓥 𝓒) 37 | (alg-fusion : 38 | ∀ {i σ Γ Δ Θ} (b : ⟦ d ⟧ (Scope (Tm d i)) σ Γ) {ρ₁ ρ₃} {ρ₂ : (Δ ─Env) 𝓥 Θ} → 39 | All Eqᴿ _ (select ρ₁ ρ₂) ρ₃ → 40 | let f = λ Δ σ → Semantics.body S ρ₂ Δ σ ∘ reify vl^Var Δ σ ∘ Semantics.body Ren ρ₁ Δ σ 41 | g = Semantics.body S ρ₃ 42 | in ⟦ d ⟧ᴿ (Kripkeᴿ Eqᴿ Eqᴿ) (fmap d f b) (fmap d g b) → 43 | Semantics.alg S (fmap d f b) ≡ Semantics.alg S (fmap d g b)) 44 | where 45 | 46 | module Ren = Semantics (Ren {d = d}) 47 | 48 | ren-sem : Fusion d Ren S S (λ Γ Δ σ → All Eqᴿ Γ ∘ (select σ)) Eqᴿ Eqᴿ 49 | Fusion.reifyᴬ ren-sem = λ _ t → t 50 | Fusion.vl^𝓥ᴬ ren-sem = vl^Var 51 | Fusion.th^𝓔ᴿ ren-sem = λ ρᴿ σ → packᴿ (λ v → cong (λ ρ → Semantics.th^𝓥 S ρ σ) (lookupᴿ ρᴿ v)) 52 | lookupᴿ (Fusion._++^Envᴿ_ ren-sem {Γ} {Δ} {ρ₁} {Ω} {ρ₂} {ρ₃} {Θ} {vs} {ws} ρᴿ vsᴿ) v with split Θ v 53 | ... | inj₁ vˡ = begin 54 | lookup (vs ++^Env ρ₂) (injectˡ Δ (lookup (base vl^Var) vˡ)) 55 | ≡⟨ injectˡ-++^Env vs ρ₂ (lookup (base vl^Var) vˡ) ⟩ 56 | lookup vs (lookup (base vl^Var) vˡ) 57 | ≡⟨ cong (lookup vs) (lookup-base^Var vˡ) ⟩ 58 | lookup vs vˡ 59 | ≡⟨ lookupᴿ vsᴿ vˡ ⟩ 60 | lookup ws vˡ 61 | ∎ 62 | ... | inj₂ vʳ = begin 63 | lookup (vs ++^Env ρ₂) (injectʳ Θ (lookup (base vl^Var) (lookup ρ₁ vʳ))) 64 | ≡⟨ injectʳ-++^Env vs ρ₂ (lookup (base vl^Var) (lookup ρ₁ vʳ)) ⟩ 65 | lookup ρ₂ (lookup (base vl^Var) (lookup ρ₁ vʳ)) 66 | ≡⟨ cong (lookup ρ₂) (lookup-base^Var (lookup ρ₁ vʳ)) ⟩ 67 | lookup ρ₂ (lookup ρ₁ vʳ) 68 | ≡⟨ lookupᴿ ρᴿ vʳ ⟩ 69 | lookup ρ₃ vʳ 70 | ∎ 71 | Fusion.varᴿ ren-sem = λ ρᴿ v → cong (Semantics.var S) (lookupᴿ ρᴿ v) 72 | Fusion.algᴿ ren-sem {Γ} {Δ} {ρ₁} {Θ} {ρ₂} {ρ₃} ρᴿ b zp = begin 73 | let 74 | v₁ = fmap d (Ren.body ρ₁) b 75 | v₃ = fmap d (Semantics.body S ρ₃) b 76 | 77 | aux : fmap d (λ Δ σ → Semantics.body S ρ₂ Δ σ ∘ reify vl^Var Δ σ) v₁ 78 | ≡ fmap d (λ Δ σ → Semantics.body S ρ₂ Δ σ ∘ reify vl^Var Δ σ ∘ Ren.body ρ₁ Δ σ) b 79 | aux = begin 80 | fmap d (λ Δ σ → Semantics.body S ρ₂ Δ σ ∘ reify vl^Var Δ σ) v₁ 81 | ≡⟨ fmap² d (Ren.body ρ₁) (λ Δ σ → Semantics.body S ρ₂ Δ σ ∘ reify vl^Var Δ σ) b ⟩ 82 | fmap d (λ Δ σ → Semantics.body S ρ₂ Δ σ ∘ reify vl^Var Δ σ ∘ Ren.body ρ₁ Δ σ) b 83 | ∎ 84 | in 85 | Semantics.alg S (fmap d (Semantics.body S ρ₂) (fmap d (reify vl^Var) v₁)) 86 | ≡⟨ cong (Semantics.alg S) (fmap² d (reify vl^Var) (Semantics.body S ρ₂) (fmap d (Ren.body ρ₁) b)) ⟩ 87 | Semantics.alg S (fmap d (λ Δ σ → Semantics.body S ρ₂ Δ σ ∘ reify vl^Var Δ σ) (fmap d (Ren.body ρ₁) b)) 88 | ≡⟨ cong (Semantics.alg S) aux ⟩ 89 | Semantics.alg S (fmap d (λ Δ σ → Semantics.body S ρ₂ Δ σ ∘ reify vl^Var Δ σ ∘ Ren.body ρ₁ Δ σ) b) 90 | ≡⟨ alg-fusion b ρᴿ (subst (λ t → ⟦ d ⟧ᴿ _ t v₃) aux zp) ⟩ 91 | Semantics.alg S v₃ 92 | ∎ 93 | \end{code} 94 | -------------------------------------------------------------------------------- /doc/generic-syntax.agda/Generic/Fusion.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | {-# OPTIONS --safe --sized-types #-} 3 | 4 | open import Data.Var hiding (z; s; _<$>_) 5 | 6 | module Generic.Fusion {I : Set} {𝓥ᴬ 𝓥ᴮ 𝓥ᴬᴮ 𝓒ᴬ 𝓒ᴮ 𝓒ᴬᴮ : I ─Scoped} where 7 | 8 | open import Size 9 | open import Data.List hiding ([_] ; zip ; lookup) 10 | open import Function renaming (_∘′_ to _∘_) hiding (_∘_) 11 | open import Relation.Binary.PropositionalEquality hiding ([_]) 12 | 13 | open import Relation.Unary 14 | open import Data.Relation hiding (_++^Envᴿ_) 15 | open import Data.Var.Varlike 16 | open import Data.Environment 17 | 18 | open import Generic.Syntax 19 | open import Generic.Semantics 20 | open import Generic.Semantics.Syntactic 21 | open import Generic.Relator 22 | 23 | private 24 | variable 25 | s : Size 26 | σ τ : I 27 | Γ Δ Θ Ω : List I 28 | ρᴬ : (Γ ─Env) 𝓥ᴬ Δ 29 | ρᴮ : (Δ ─Env) 𝓥ᴮ Θ 30 | ρᴬᴮ : (Γ ─Env) 𝓥ᴬᴮ Θ 31 | vsᴬᴮ : (Δ ─Env) 𝓥ᴬᴮ Γ 32 | vsᴮ : (Δ ─Env) 𝓥ᴮ Γ 33 | 34 | \end{code} 35 | %<*fusionrec> 36 | \begin{code} 37 | record Fusion (d : Desc I) (𝓢ᴬ : Semantics d 𝓥ᴬ 𝓒ᴬ) (𝓢ᴮ : Semantics d 𝓥ᴮ 𝓒ᴮ) 38 | (𝓢ᴬᴮ : Semantics d 𝓥ᴬᴮ 𝓒ᴬᴮ) 39 | (𝓔ᴿ : ∀ Γ Δ {Θ} → (Γ ─Env) 𝓥ᴬ Δ → (Δ ─Env) 𝓥ᴮ Θ → (Γ ─Env) 𝓥ᴬᴮ Θ → Set) 40 | (𝓥ᴿ : Rel 𝓥ᴮ 𝓥ᴬᴮ) (𝓒ᴿ : Rel 𝓒ᴮ 𝓒ᴬᴮ) : Set where 41 | \end{code} 42 | % 43 | \begin{code} 44 | module 𝓢ᴬ = Semantics 𝓢ᴬ 45 | module 𝓢ᴮ = Semantics 𝓢ᴮ 46 | module 𝓢ᴬᴮ = Semantics 𝓢ᴬᴮ 47 | evalᴬ = 𝓢ᴬ.semantics 48 | evalᴮ = 𝓢ᴮ.semantics 49 | evalᴬᴮ = 𝓢ᴬᴮ.semantics 50 | field 51 | \end{code} 52 | %<*reify> 53 | \begin{code} 54 | reifyᴬ : ∀ σ → ∀[ 𝓒ᴬ σ ⇒ Tm d ∞ σ ] 55 | \end{code} 56 | % 57 | %<*vlV> 58 | \begin{code} 59 | vl^𝓥ᴬ : VarLike 𝓥ᴬ 60 | \end{code} 61 | % 62 | %<*quote> 63 | \begin{code} 64 | quoteᴬ : ∀ Δ i → Kripke 𝓥ᴬ 𝓒ᴬ Δ i Γ → Tm d ∞ i (Δ ++ Γ) 65 | quoteᴬ Δ i k = reifyᴬ i (reify vl^𝓥ᴬ Δ i k) 66 | \end{code} 67 | % 68 | \begin{code} 69 | field 70 | \end{code} 71 | %<*appendR> 72 | \begin{code} 73 | _++^Envᴿ_ : 𝓔ᴿ Γ Δ ρᴬ ρᴮ ρᴬᴮ → All 𝓥ᴿ Θ vsᴮ vsᴬᴮ → 74 | let id++^Envρᴬ = freshˡ vl^𝓥ᴬ Δ ++^Env th^Env 𝓢ᴬ.th^𝓥 ρᴬ (freshʳ vl^Var Θ) 75 | in 𝓔ᴿ (Θ ++ Γ) (Θ ++ Δ) id++^Envρᴬ (vsᴮ ++^Env ρᴮ) (vsᴬᴮ ++^Env ρᴬᴮ) 76 | \end{code} 77 | % 78 | %<*thV> 79 | \begin{code} 80 | th^𝓔ᴿ : 𝓔ᴿ Γ Δ ρᴬ ρᴮ ρᴬᴮ → (ρ : Thinning Θ Ω) → 81 | 𝓔ᴿ Γ Δ ρᴬ (th^Env 𝓢ᴮ.th^𝓥 ρᴮ ρ) (th^Env 𝓢ᴬᴮ.th^𝓥 ρᴬᴮ ρ) 82 | \end{code} 83 | % 84 | %<*crel> 85 | \begin{code} 86 | 𝓡 : ∀ σ → (Γ ─Env) 𝓥ᴬ Δ → (Δ ─Env) 𝓥ᴮ Θ → (Γ ─Env) 𝓥ᴬᴮ Θ → 87 | Tm d s σ Γ → Set 88 | 𝓡 σ ρᴬ ρᴮ ρᴬᴮ t = rel 𝓒ᴿ σ (evalᴮ ρᴮ (reifyᴬ σ (evalᴬ ρᴬ t))) (evalᴬᴮ ρᴬᴮ t) 89 | \end{code} 90 | % 91 | \begin{code} 92 | field 93 | \end{code} 94 | %<*varR> 95 | \begin{code} 96 | varᴿ : 𝓔ᴿ Γ Δ ρᴬ ρᴮ ρᴬᴮ → ∀ v → 𝓡 σ ρᴬ ρᴮ ρᴬᴮ (`var v) 97 | \end{code} 98 | % 99 | %<*algR> 100 | \begin{code} 101 | algᴿ : 𝓔ᴿ Γ Δ ρᴬ ρᴮ ρᴬᴮ → (b : ⟦ d ⟧ (Scope (Tm d s)) σ Γ) → 102 | let bᴬ : ⟦ d ⟧ (Kripke 𝓥ᴬ 𝓒ᴬ) _ _ 103 | bᴬ = fmap d (𝓢ᴬ.body ρᴬ) b 104 | bᴮ = fmap d (λ Δ i → 𝓢ᴮ.body ρᴮ Δ i ∘ quoteᴬ Δ i) bᴬ 105 | bᴬᴮ = fmap d (𝓢ᴬᴮ.body ρᴬᴮ) b 106 | in ⟦ d ⟧ᴿ (Kripkeᴿ 𝓥ᴿ 𝓒ᴿ) bᴮ bᴬᴮ → 𝓡 σ ρᴬ ρᴮ ρᴬᴮ (`con b) 107 | \end{code} 108 | % 109 | %<*fusiontype> 110 | \begin{code} 111 | fusion : 𝓔ᴿ Γ Δ ρᴬ ρᴮ ρᴬᴮ → (t : Tm d s σ Γ) → 𝓡 σ ρᴬ ρᴮ ρᴬᴮ t 112 | \end{code} 113 | % 114 | %<*bodytype> 115 | \begin{code} 116 | body : 𝓔ᴿ Γ Δ ρᴬ ρᴮ ρᴬᴮ → ∀ Δ σ → (b : Scope (Tm d s) Δ σ Γ) → 117 | let vᴮ = 𝓢ᴮ.body ρᴮ Δ σ (quoteᴬ Δ σ (𝓢ᴬ.body ρᴬ Δ σ b)) 118 | vᴬᴮ = 𝓢ᴬᴮ.body ρᴬᴮ Δ σ b 119 | in Kripkeᴿ 𝓥ᴿ 𝓒ᴿ Δ σ vᴮ vᴬᴮ 120 | \end{code} 121 | % 122 | %<*fusioncode> 123 | \begin{code} 124 | fusion ρᴿ (`var v) = varᴿ ρᴿ v 125 | fusion ρᴿ (`con t) = algᴿ ρᴿ t (rew (liftᴿ d (body ρᴿ) t)) where 126 | 127 | eq = fmap² d (𝓢ᴬ.body _) (λ Δ i t → 𝓢ᴮ.body _ Δ i (quoteᴬ Δ i t)) t 128 | rew = subst (λ v → ⟦ d ⟧ᴿ (Kripkeᴿ 𝓥ᴿ 𝓒ᴿ) v _) (sym eq) 129 | \end{code} 130 | % 131 | %<*bodycode> 132 | \begin{code} 133 | body ρᴿ [] i b = fusion ρᴿ b 134 | body ρᴿ (σ ∷ Δ) i b = λ ρ vsᴿ → fusion (th^𝓔ᴿ ρᴿ ρ ++^Envᴿ vsᴿ) b 135 | \end{code} 136 | % 137 | -------------------------------------------------------------------------------- /doc/shared/Data/Var/Varlike.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | {-# OPTIONS --safe --sized-types #-} 3 | 4 | module Data.Var.Varlike where 5 | 6 | open import Data.List.Base hiding (lookup ; [_]) 7 | open import Function.Base 8 | open import Relation.Binary.PropositionalEquality hiding ([_]) 9 | 10 | open import Relation.Unary using (IUniversal; _⊢_; _⇒_) 11 | open import Data.Var 12 | open import Data.Pred as Pred hiding (All) 13 | open import Data.Relation 14 | open import Data.Environment hiding (uncurry) 15 | open import Generic.Syntax 16 | 17 | private 18 | variable 19 | I : Set 20 | σ : I 21 | Γ Δ : List I 22 | 𝓥 𝓥₁ 𝓥₂ 𝓒 𝓥ᴬ 𝓥ᴮ 𝓒ᴬ 𝓒ᴮ : I ─Scoped 23 | 24 | \end{code} 25 | %<*varlike> 26 | \begin{code} 27 | record VarLike (𝓥 : I ─Scoped) : Set where 28 | field th^𝓥 : Thinnable (𝓥 σ) 29 | new : ∀[ (σ ∷_) ⊢ 𝓥 σ ] 30 | \end{code} 31 | % 32 | \begin{code} 33 | 34 | base : (Γ ─Env) 𝓥 Γ 35 | base {Γ = []} = ε 36 | base {Γ = σ ∷ Γ} = th^Env th^𝓥 base extend ∙ new 37 | 38 | freshʳ : (Δ : List I) → (Γ ─Env) 𝓥 (Δ ++ Γ) 39 | freshʳ Δ = th^Env th^𝓥 base (pack (injectʳ Δ)) 40 | 41 | freshˡ : (Δ : List I) → (Γ ─Env) 𝓥 (Γ ++ Δ) 42 | freshˡ k = th^Env th^𝓥 base (pack (injectˡ _)) 43 | 44 | singleton : 𝓥 σ Γ → (σ ∷ Γ ─Env) 𝓥 Γ 45 | singleton v = base ∙ v 46 | open VarLike public 47 | 48 | vl^Var : VarLike {I} Var 49 | new vl^Var = z 50 | th^𝓥 vl^Var = th^Var 51 | 52 | lookup-base^Var : (k : Var σ Γ) → lookup (base vl^Var) k ≡ k 53 | lookup-base^Var z = refl 54 | lookup-base^Var (s k) = cong s (lookup-base^Var k) 55 | 56 | \end{code} 57 | %<*reify> 58 | \begin{code} 59 | reify : VarLike 𝓥 → ∀ Δ i → Kripke 𝓥 𝓒 Δ i Γ → Scope 𝓒 Δ i Γ 60 | reify vl^𝓥 [] i b = b 61 | reify vl^𝓥 Δ@(_ ∷ _) i b = b (freshʳ vl^Var Δ) (freshˡ vl^𝓥 _) 62 | \end{code} 63 | % 64 | \begin{code} 65 | module _ (vl^𝓥 : VarLike 𝓥) where 66 | 67 | lift : ∀ Θ → (Γ ─Env) 𝓥 Δ → ((Θ ++ Γ) ─Env) 𝓥 (Θ ++ Δ) 68 | lift Θ ρ = freshˡ vl^𝓥 _ ++^Env th^Env (th^𝓥 vl^𝓥) ρ (freshʳ vl^Var Θ) 69 | 70 | extend-is-fresh : All Eqᴿ Γ extend (freshʳ vl^Var (σ ∷ [])) 71 | lookupᴿ extend-is-fresh k = cong s (sym (lookup-base^Var k)) 72 | 73 | module _ {I : Set} {𝓥 : I ─Scoped} where 74 | open ≡-Reasoning 75 | 76 | freshʳ-++^Env : (Δ : List I) {Γ Θ : List I} 77 | (ρ₁ : (Δ ─Env) 𝓥 Θ) (ρ₂ : (Γ ─Env) 𝓥 Θ) {i : I} (v : Var i Γ) → 78 | lookup (ρ₁ ++^Env ρ₂) (lookup (freshʳ vl^Var Δ) v) ≡ lookup ρ₂ v 79 | freshʳ-++^Env Δ ρ₁ ρ₂ v = begin 80 | lookup (ρ₁ ++^Env ρ₂) (lookup (freshʳ vl^Var Δ) v) 81 | ≡⟨ injectʳ-++^Env ρ₁ ρ₂ (lookup (base vl^Var) v) ⟩ 82 | lookup ρ₂ (lookup (base vl^Var) v) 83 | ≡⟨ cong (lookup ρ₂) (lookup-base^Var v) ⟩ 84 | lookup ρ₂ v 85 | ∎ 86 | 87 | module _ (𝓡^𝓥 : Rel 𝓥₁ 𝓥₂) where 88 | 89 | record VarLikeᴿ (vl₁ : VarLike 𝓥₁) (vl₂ : VarLike 𝓥₂) : Set where 90 | field newᴿ : rel 𝓡^𝓥 σ {σ ∷ Γ} (new vl₁) (new vl₂) 91 | thᴿ : (ρ : Thinning Γ Δ) {v₁ : 𝓥₁ σ Γ} {v₂ : 𝓥₂ σ Γ} → 92 | rel 𝓡^𝓥 σ v₁ v₂ → rel 𝓡^𝓥 σ (th^𝓥 vl₁ v₁ ρ) (th^𝓥 vl₂ v₂ ρ) 93 | 94 | baseᴿ : All 𝓡^𝓥 Γ (base vl₁) (base vl₂) 95 | baseᴿ {[] } = packᴿ λ () 96 | baseᴿ {i ∷ Γ} = (thᴿ extend <$>ᴿ baseᴿ) ∙ᴿ newᴿ 97 | 98 | freshˡᴿ : ∀ Γ → All 𝓡^𝓥 Δ (freshˡ vl₁ Γ) (freshˡ vl₂ Γ) 99 | freshˡᴿ n = thᴿ _ <$>ᴿ baseᴿ 100 | 101 | freshʳᴿ : ∀ Γ → All 𝓡^𝓥 Δ (freshʳ vl₁ Γ) (freshʳ vl₂ Γ) 102 | freshʳᴿ n = thᴿ _ <$>ᴿ baseᴿ 103 | 104 | 105 | module _ (vl^𝓥 : VarLike 𝓥) where 106 | 107 | vl^Refl : VarLikeᴿ Eqᴿ vl^𝓥 vl^𝓥 108 | VarLikeᴿ.newᴿ vl^Refl = refl 109 | VarLikeᴿ.thᴿ vl^Refl = λ σ → cong (λ v → th^𝓥 vl^𝓥 v σ) 110 | 111 | 112 | module _ (𝓥ᴾ : Pred {I} 𝓥) (𝓒ᴾ : Pred {I} 𝓒) where 113 | 114 | Kripkeᴾ : ∀ Δ τ → ∀[ Kripke 𝓥 𝓒 Δ τ ⇒ const Set ] 115 | Kripkeᴾ [] τ k = pred 𝓒ᴾ τ k 116 | Kripkeᴾ Δ@(_ ∷ _) τ k = ∀ {Θ} th {ρ : (Δ ─Env) 𝓥 Θ} → Pred.All 𝓥ᴾ Δ ρ → pred 𝓒ᴾ τ (k th ρ) 117 | 118 | module _ (𝓥ᴿ : Rel {I} 𝓥ᴬ 𝓥ᴮ) (𝓒ᴿ : Rel {I} 𝓒ᴬ 𝓒ᴮ) where 119 | \end{code} 120 | %<*kripkeR> 121 | \begin{code} 122 | Kripkeᴿ : ∀ Δ i → ∀[ Kripke 𝓥ᴬ 𝓒ᴬ Δ i ⇒ Kripke 𝓥ᴮ 𝓒ᴮ Δ i ⇒ const Set ] 123 | Kripkeᴿ [] σ kᴬ kᴮ = rel 𝓒ᴿ σ kᴬ kᴮ 124 | Kripkeᴿ Δ@(_ ∷ _) σ kᴬ kᴮ = ∀ {Θ} (ρ : Thinning _ Θ) {vsᴬ vsᴮ} → 125 | All 𝓥ᴿ Δ vsᴬ vsᴮ → rel 𝓒ᴿ σ (kᴬ ρ vsᴬ) (kᴮ ρ vsᴮ) 126 | \end{code} 127 | % 128 | -------------------------------------------------------------------------------- /doc/generic-syntax.agda/Generic/Fusion/Utils.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | {-# OPTIONS --safe --sized-types #-} 3 | 4 | module Generic.Fusion.Utils where 5 | 6 | open import Data.Var hiding (_<$>_) 7 | 8 | open import Size 9 | open import Data.List hiding ([_] ; zip ; lookup) 10 | open import Function renaming (_∘′_ to _∘_) hiding (_∘_) 11 | open import Relation.Unary 12 | open import Relation.Binary.PropositionalEquality hiding ([_]) 13 | 14 | open import Data.Relation hiding (_++^Envᴿ_) 15 | open import Data.Var.Varlike 16 | open import Data.Environment 17 | 18 | open import Generic.Syntax 19 | open import Generic.Semantics 20 | open import Generic.Semantics.Syntactic 21 | open import Generic.Fusion 22 | 23 | module _ {I : Set} {T : I ─Scoped} where 24 | 25 | open ≡-Reasoning 26 | 27 | -- this is the shape of environment one obtains when pushing an evaluation environment 28 | -- on top of a thinning into the body of a binder 29 | 30 | thBodyEnv : 31 | ∀ {Γ Δ Θ Ξ : List I} {ρᴬ : Thinning Γ Δ} {ρᴮ : (Δ ─Env) T Θ} 32 | {ρᴬᴮ : (Γ ─Env) T Θ} {ρ₄ ρ₅ : (Ξ ─Env) T Θ} 33 | (ρᴿ : All Eqᴿ Γ (select ρᴬ ρᴮ) ρᴬᴮ) (vsᴿ : All Eqᴿ Ξ ρ₄ ρ₅) → 34 | let σ : (Ξ ++ Γ ─Env) Var (Ξ ++ Δ) 35 | σ = freshˡ vl^Var Δ ++^Env th^Env th^Var ρᴬ (freshʳ vl^Var Ξ) 36 | in All Eqᴿ (Ξ ++ Γ) (select σ (ρ₄ ++^Env ρᴮ)) (ρ₅ ++^Env ρᴬᴮ) 37 | lookupᴿ (thBodyEnv {Γ} {Δ} {Θ} {Ξ} {ρᴬ} {ρᴮ} {ρᴬᴮ} {ρ₄} {ρ₅} ρᴿ vsᴿ) k 38 | with split Ξ k 39 | ... | inj₁ kˡ = begin 40 | lookup (ρ₄ ++^Env ρᴮ) (injectˡ Δ (lookup (base vl^Var) kˡ)) 41 | ≡⟨ injectˡ-++^Env ρ₄ ρᴮ (lookup (base vl^Var) kˡ) ⟩ 42 | lookup ρ₄ (lookup (base vl^Var) kˡ) 43 | ≡⟨ cong (lookup ρ₄) (lookup-base^Var kˡ) ⟩ 44 | lookup ρ₄ kˡ 45 | ≡⟨ lookupᴿ vsᴿ kˡ ⟩ 46 | lookup ρ₅ kˡ 47 | ∎ 48 | ... | inj₂ kʳ = begin 49 | lookup (ρ₄ ++^Env ρᴮ) (injectʳ Ξ (lookup (base vl^Var) (lookup ρᴬ kʳ))) 50 | ≡⟨ injectʳ-++^Env ρ₄ ρᴮ (lookup (base vl^Var) (lookup ρᴬ kʳ)) ⟩ 51 | lookup ρᴮ (lookup (base vl^Var) (lookup ρᴬ kʳ)) 52 | ≡⟨ cong (lookup ρᴮ) (lookup-base^Var (lookup ρᴬ kʳ)) ⟩ 53 | lookup ρᴮ (lookup ρᴬ kʳ) 54 | ≡⟨ lookupᴿ ρᴿ kʳ ⟩ 55 | lookup ρᴬᴮ kʳ 56 | ∎ 57 | 58 | module _ {I : Set} {d : Desc I} {𝓥 𝓒 : I ─Scoped} 59 | (𝓢 : Semantics d 𝓥 𝓒) 60 | (𝓕 : Fusion d Ren 𝓢 𝓢 (λ Γ Δ ρᴬ ρᴮ → All Eqᴿ _ (select ρᴬ ρᴮ)) Eqᴿ Eqᴿ) 61 | (eq^quote : ∀ σ {Γ} (t : Tm d ∞ σ Γ) → Fusion.reifyᴬ 𝓕 σ t ≡ t) where 62 | 63 | open ≡-Reasoning 64 | module 𝓢 = Semantics 𝓢 65 | 66 | SemVarTmᴿ : Rel 𝓥 𝓒 67 | rel SemVarTmᴿ σ v c = 𝓢.var v ≡ c 68 | 69 | -- this is the shape of environment one obtains when pushing an evaluation environment 70 | -- on top of a substitution into the body of a binder 71 | 72 | subBodyEnv : 73 | ∀ {Γ Δ Θ Ξ} (ρᴬ : (Γ ─Env) (Tm d _) Δ) {ρᴮ : (Δ ─Env) 𝓥 Θ} {ρᴬᴮ} 74 | {ρ₄ : (Ξ ─Env) 𝓥 Θ} {ρ₅ : (Ξ ─Env) 𝓒 Θ} → 75 | All Eqᴿ Γ (𝓢.semantics ρᴮ <$> ρᴬ) ρᴬᴮ → 76 | All SemVarTmᴿ _ ρ₄ ρ₅ → 77 | let σ : ((Ξ ++ Γ) ─Env) (Tm d _) (Ξ ++ Δ) 78 | σ = freshˡ vl^Tm Δ ++^Env th^Env th^Tm ρᴬ (freshʳ vl^Var Ξ) 79 | in All Eqᴿ (Ξ ++ Γ) (𝓢.semantics (ρ₄ ++^Env ρᴮ) <$> σ) (ρ₅ ++^Env ρᴬᴮ) 80 | lookupᴿ (subBodyEnv {Γ} {Δ} {Θ} {Ξ} ρᴬ {ρᴮ} {ρᴬᴮ} {ρ₄} {ρ₅} ρᴿ vsᴿ) k 81 | with split Ξ k 82 | ... | inj₁ kˡ = begin 83 | let t = ren (pack (injectˡ Δ)) (lookup (base vl^Tm) kˡ) in 84 | 𝓢.semantics (ρ₄ ++^Env ρᴮ) t 85 | ≡⟨ cong (𝓢.semantics (ρ₄ ++^Env ρᴮ)) (sym (eq^quote _ t)) ⟩ 86 | 𝓢.semantics (ρ₄ ++^Env ρᴮ) (Fusion.reifyᴬ 𝓕 _ t) 87 | ≡⟨ Fusion.fusion 𝓕 (packᴿ (injectˡ-++^Env ρ₄ ρᴮ)) (lookup (base vl^Tm) kˡ) ⟩ 88 | 𝓢.semantics ρ₄ (lookup (base vl^Tm) kˡ) 89 | ≡⟨ cong (𝓢.semantics ρ₄) (lookup-base^Tm kˡ) ⟩ 90 | 𝓢.var(lookup ρ₄ kˡ) 91 | ≡⟨ lookupᴿ vsᴿ kˡ ⟩ 92 | lookup ρ₅ kˡ 93 | ∎ 94 | ... | inj₂ kʳ = begin 95 | let t = ren (freshʳ vl^Var Ξ) (lookup ρᴬ kʳ) in 96 | 𝓢.semantics (ρ₄ ++^Env ρᴮ) t 97 | ≡⟨ cong (𝓢.semantics (ρ₄ ++^Env ρᴮ)) (sym (eq^quote _ t)) ⟩ 98 | 𝓢.semantics (ρ₄ ++^Env ρᴮ) (Fusion.reifyᴬ 𝓕 _ t) 99 | ≡⟨ Fusion.fusion 𝓕 eqᴿ (lookup ρᴬ kʳ) ⟩ 100 | 𝓢.semantics ρᴮ (lookup ρᴬ kʳ) 101 | ≡⟨ lookupᴿ ρᴿ kʳ ⟩ 102 | lookup ρᴬᴮ kʳ 103 | ∎ where 104 | 105 | eqᴿ : All Eqᴿ Δ (select (freshʳ vl^Var Ξ) (ρ₄ ++^Env ρᴮ)) ρᴮ 106 | lookupᴿ eqᴿ v = begin 107 | lookup (select (freshʳ vl^Var Ξ) (ρ₄ ++^Env ρᴮ)) v 108 | ≡⟨ injectʳ-++^Env ρ₄ ρᴮ (lookup (base vl^Var) v) ⟩ 109 | lookup ρᴮ (lookup (base vl^Var) v) 110 | ≡⟨ cong (lookup ρᴮ) (lookup-base^Var v) ⟩ 111 | lookup ρᴮ v 112 | ∎ 113 | \end{code} 114 | -------------------------------------------------------------------------------- /doc/type-scope-semantics.agda/Properties/Fusion/Syntactic/Instances.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | module Properties.Fusion.Syntactic.Instances where 3 | 4 | open import Syntax.Calculus 5 | open import Syntax.Type 6 | open import Data.Var hiding (_<$>_) 7 | open import Data.List.Base using (List; []; _∷_) 8 | open import Data.Environment 9 | open import Semantics.Syntactic.Instances 10 | open import Data.Relation as Rel hiding (_∙ᴿ_) 11 | open import Properties.Fusion.Syntactic.Specification as Syntactic 12 | open import Properties.Fusion.Specification as Fusion 13 | open import Properties.Simulation.Instances using (VarTermᴿ) 14 | 15 | open import Function 16 | open import Relation.Binary.PropositionalEquality hiding (trans) 17 | open ≡-Reasoning 18 | 19 | open SynFusion 20 | 21 | 22 | \end{code} 23 | %<*renrenfus> 24 | \begin{code} 25 | RenRen : SynFusion Syn^Ren Syn^Ren Syn^Ren 26 | (λ ρᴬ ρᴮ → All Eqᴿ _ (select ρᴬ ρᴮ)) Eqᴿ 27 | RenRen ._∙ᴿ_ = λ ρᴿ tᴿ → packᴿ λ where 28 | z → tᴿ 29 | (s v) → lookupᴿ ρᴿ v 30 | RenRen .th^𝓔ᴿ = λ ρᴿ ρ → cong (λ v → th^Var v ρ) <$>ᴿ ρᴿ 31 | RenRen .varᴿ = λ ρᴿ v → cong `var (lookupᴿ ρᴿ v) 32 | RenRen .zroᴿ = refl 33 | \end{code} 34 | % 35 | \begin{code} 36 | private 37 | variable 38 | σ : Type 39 | Γ Δ : List Type 40 | 𝓣 : Type ─Scoped 41 | ρ₁ ρ₂ : (Γ ─Env) 𝓣 Δ 42 | t : Term σ Γ 43 | 44 | \end{code} 45 | %<*renren> 46 | \begin{code} 47 | renren : (t : Term σ Γ) → ren ρ₂ (ren ρ₁ t) ≡ ren (select ρ₁ ρ₂) t 48 | renren = let fus = Syntactic.Fundamental.lemma RenRen 49 | in Fusion.fusion fus reflᴿ 50 | \end{code} 51 | % 52 | %<*rensubfus> 53 | \begin{code} 54 | RenSub : SynFusion Syn^Ren Syn^Sub Syn^Sub 55 | (λ ρᴬ ρᴮ → All Eqᴿ _ (select ρᴬ ρᴮ)) Eqᴿ 56 | \end{code} 57 | % 58 | \begin{code} 59 | RenSub ._∙ᴿ_ = λ ρᴿ tᴿ → packᴿ λ { z → tᴿ ; (s v) → lookupᴿ ρᴿ v } 60 | RenSub .th^𝓔ᴿ = λ ρᴿ ρ → cong (λ t → th^Term t ρ) <$>ᴿ ρᴿ 61 | RenSub .varᴿ = λ ρᴿ v → lookupᴿ ρᴿ v 62 | RenSub .zroᴿ = refl 63 | 64 | \end{code} 65 | %<*rensub> 66 | \begin{code} 67 | rensub : (t : Term σ Γ) → sub ρ₂ (ren ρ₁ t) ≡ sub (select ρ₁ ρ₂) t 68 | rensub = let fus = Syntactic.Fundamental.lemma RenSub 69 | in Fusion.fusion fus reflᴿ 70 | \end{code} 71 | % 72 | %<*subrenfus> 73 | \begin{code} 74 | SubRen : SynFusion Syn^Sub Syn^Ren Syn^Sub 75 | (λ ρᴬ ρᴮ → All Eqᴿ _ (ren ρᴮ <$> ρᴬ)) VarTermᴿ 76 | \end{code} 77 | % 78 | \begin{code} 79 | SubRen ._∙ᴿ_ {ρᴬ = ρᴬ} {ρᴮ = ρᴮ} {ρᴬᴮ = ρᴬᴮ} {tᴮ = tᴮ} {tᴬᴮ = tᴬᴮ} = λ ρᴿ tᴿ → packᴿ λ where 80 | z → tᴿ 81 | (s v) → begin 82 | th^Term (th^Term (lookup ρᴬ v) extend) (ρᴮ ∙ tᴮ) 83 | ≡⟨ renren (lookup ρᴬ v) ⟩ 84 | th^Term (lookup ρᴬ v) ρᴮ 85 | ≡⟨ lookupᴿ ρᴿ v ⟩ 86 | lookup ρᴬᴮ v 87 | ∎ 88 | SubRen .th^𝓔ᴿ {ρᴬ = ρᴬ} {ρᴮ = ρᴮ} {ρᴬᴮ = ρᴬᴮ} = λ ρᴿ ρ → packᴿ λ v → begin 89 | th^Term (lookup ρᴬ v) (select ρᴮ ρ) 90 | ≡⟨ sym (renren (lookup ρᴬ v)) ⟩ 91 | th^Term (th^Term (lookup ρᴬ v) ρᴮ) ρ 92 | ≡⟨ cong (λ t → th^Term t ρ) (lookupᴿ ρᴿ v) ⟩ 93 | th^Term (lookup ρᴬᴮ v) ρ 94 | ∎ 95 | SubRen .varᴿ = λ ρᴿ v → lookupᴿ ρᴿ v 96 | SubRen .zroᴿ = refl 97 | 98 | \end{code} 99 | %<*subren> 100 | \begin{code} 101 | subren : (t : Term σ Γ) → ren ρ₂ (sub ρ₁ t) ≡ sub (ren ρ₂ <$> ρ₁) t 102 | subren = let fus = Syntactic.Fundamental.lemma SubRen 103 | in Fusion.fusion fus reflᴿ 104 | \end{code} 105 | % 106 | %<*subsubfus> 107 | \begin{code} 108 | SubSub : SynFusion Syn^Sub Syn^Sub Syn^Sub 109 | (λ ρᴬ ρᴮ → All Eqᴿ _ (sub ρᴮ <$> ρᴬ)) Eqᴿ 110 | \end{code} 111 | % 112 | \begin{code} 113 | SubSub ._∙ᴿ_ {ρᴬ = ρᴬ} {ρᴮ = ρᴮ} {ρᴬᴮ = ρᴬᴮ} {tᴮ = tᴮ} {tᴬᴮ = tᴬᴮ} = λ ρᴿ tᴿ → packᴿ λ where 114 | z → tᴿ 115 | (s v) → begin 116 | sub (ρᴮ ∙ tᴮ) (th^Term (lookup ρᴬ v) extend) 117 | ≡⟨ rensub (lookup ρᴬ v) ⟩ 118 | sub ρᴮ (lookup ρᴬ v) 119 | ≡⟨ lookupᴿ ρᴿ v ⟩ 120 | lookup ρᴬᴮ v 121 | ∎ 122 | SubSub .th^𝓔ᴿ {ρᴬ = ρᴬ} {ρᴮ = ρᴮ} {ρᴬᴮ = ρᴬᴮ} = λ ρᴿ ρ → packᴿ λ v → begin 123 | sub ((λ t → th^Term t ρ) <$> ρᴮ) (lookup ρᴬ v) 124 | ≡⟨ sym (subren (lookup ρᴬ v)) ⟩ 125 | th^Term (sub ρᴮ (lookup ρᴬ v)) ρ 126 | ≡⟨ cong (λ t → th^Term t ρ) (lookupᴿ ρᴿ v) ⟩ 127 | th^Term (lookup ρᴬᴮ v) ρ 128 | ∎ 129 | SubSub .varᴿ = λ ρᴿ v → lookupᴿ ρᴿ v 130 | SubSub .zroᴿ = refl 131 | 132 | \end{code} 133 | %<*subsub> 134 | \begin{code} 135 | subsub : (t : Term σ Γ) → sub ρ₁ (sub ρ₂ t) ≡ sub (sub ρ₁ <$> ρ₂) t 136 | subsub = let fus = Syntactic.Fundamental.lemma SubSub 137 | in Fusion.fusion fus reflᴿ 138 | \end{code} 139 | % 140 | -------------------------------------------------------------------------------- /doc/type-scope-semantics.agda/Properties/Fusion/Specification.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | open import Syntax.Type 3 | open import Data.Var 4 | 5 | module Properties.Fusion.Specification {𝓥ᴬ 𝓒ᴬ 𝓥ᴮ 𝓒ᴮ 𝓥ᴬᴮ 𝓒ᴬᴮ : Type ─Scoped} where 6 | 7 | open import Data.Environment 8 | open import Data.List.Base using (List; []; _∷_) 9 | open import Data.Relation hiding (_∙ᴿ_) 10 | open import Syntax.Type 11 | open import Syntax.Calculus 12 | open import Semantics.Specification 13 | open import Function renaming (_$′_ to _$_) using () 14 | open import Relation.Unary 15 | 16 | private 17 | variable 18 | σ τ : Type 19 | Γ Δ Θ Ω : List Type 20 | ρᴬ : (Γ ─Env) 𝓥ᴬ Δ 21 | ρᴮ : (Δ ─Env) 𝓥ᴮ Θ 22 | ρᴬᴮ : (Γ ─Env) 𝓥ᴬᴮ Θ 23 | vᴬᴮ : 𝓥ᴬᴮ σ Γ 24 | vᴮ : 𝓥ᴮ σ Γ 25 | 26 | \end{code} 27 | %<*fusion> 28 | \begin{code} 29 | record Fusion 30 | (𝓢ᴬ : Semantics 𝓥ᴬ 𝓒ᴬ) (𝓢ᴮ : Semantics 𝓥ᴮ 𝓒ᴮ) (𝓢ᴬᴮ : Semantics 𝓥ᴬᴮ 𝓒ᴬᴮ) 31 | (𝓔ᴿ : ∀ {Γ Δ Θ} → (Γ ─Env) 𝓥ᴬ Δ → (Δ ─Env) 𝓥ᴮ Θ → (Γ ─Env) 𝓥ᴬᴮ Θ → Set) 32 | (𝓥ᴿ : Rel 𝓥ᴮ 𝓥ᴬᴮ) (𝓒ᴿ : Rel 𝓒ᴮ 𝓒ᴬᴮ) : Set where 33 | \end{code} 34 | % 35 | \begin{code} 36 | module 𝓢ᴬ = Semantics 𝓢ᴬ 37 | module 𝓢ᴮ = Semantics 𝓢ᴮ 38 | module 𝓢ᴬᴮ = Semantics 𝓢ᴬᴮ 39 | 40 | field 41 | \end{code} 42 | %<*reify> 43 | \begin{code} 44 | reifyᴬ : ∀[ 𝓒ᴬ σ ⇒ Term σ ] 45 | \end{code} 46 | % 47 | %<*var0> 48 | \begin{code} 49 | var0ᴬ : ∀[ (σ ∷_) ⊢ 𝓥ᴬ σ ] 50 | \end{code} 51 | % 52 | %<*thV> 53 | \begin{code} 54 | _∙ᴿ_ : 𝓔ᴿ ρᴬ ρᴮ ρᴬᴮ → rel 𝓥ᴿ σ vᴮ vᴬᴮ → 55 | 𝓔ᴿ (th^Env 𝓢ᴬ.th^𝓥 ρᴬ extend ∙ var0ᴬ) (ρᴮ ∙ vᴮ) (ρᴬᴮ ∙ vᴬᴮ) 56 | th^𝓔ᴿ : 𝓔ᴿ ρᴬ ρᴮ ρᴬᴮ → (ρ : Thinning Θ Ω) → 57 | 𝓔ᴿ ρᴬ (th^Env 𝓢ᴮ.th^𝓥 ρᴮ ρ) (th^Env 𝓢ᴬᴮ.th^𝓥 ρᴬᴮ ρ) 58 | \end{code} 59 | % 60 | %<*crel> 61 | \begin{code} 62 | 𝓡 : ∀ σ → (Γ ─Env) 𝓥ᴬ Δ → (Δ ─Env) 𝓥ᴮ Θ → (Γ ─Env) 𝓥ᴬᴮ Θ → 63 | Term σ Γ → Set 64 | 𝓡 σ ρᴬ ρᴮ ρᴬᴮ t = let vᴬ = semantics 𝓢ᴬ ρᴬ t 65 | vᴮ = semantics 𝓢ᴮ ρᴮ (reifyᴬ vᴬ) 66 | vᴬᴮ = semantics 𝓢ᴬᴮ ρᴬᴮ t 67 | in rel 𝓒ᴿ σ vᴮ vᴬᴮ 68 | \end{code} 69 | % 70 | \begin{code} 71 | field 72 | \end{code} 73 | %<*var> 74 | \begin{code} 75 | varᴿ : 𝓔ᴿ ρᴬ ρᴮ ρᴬᴮ → (v : Var σ Γ) → 𝓡 σ ρᴬ ρᴮ ρᴬᴮ (`var v) 76 | \end{code} 77 | % 78 | %<*base> 79 | \begin{code} 80 | oneᴿ : 𝓔ᴿ ρᴬ ρᴮ ρᴬᴮ → 𝓡 `Unit ρᴬ ρᴮ ρᴬᴮ `one 81 | ttᴿ : 𝓔ᴿ ρᴬ ρᴮ ρᴬᴮ → 𝓡 `Bool ρᴬ ρᴮ ρᴬᴮ `tt 82 | ffᴿ : 𝓔ᴿ ρᴬ ρᴮ ρᴬᴮ → 𝓡 `Bool ρᴬ ρᴮ ρᴬᴮ `ff 83 | \end{code} 84 | % 85 | %<*struct> 86 | \begin{code} 87 | appᴿ : 𝓔ᴿ ρᴬ ρᴮ ρᴬᴮ → 88 | ∀ f t → 𝓡 (σ `→ τ) ρᴬ ρᴮ ρᴬᴮ f → 𝓡 σ ρᴬ ρᴮ ρᴬᴮ t → 89 | 𝓡 τ ρᴬ ρᴮ ρᴬᴮ (`app f t) 90 | ifteᴿ : 𝓔ᴿ ρᴬ ρᴮ ρᴬᴮ → 91 | ∀ b l r → 𝓡 `Bool ρᴬ ρᴮ ρᴬᴮ b → 𝓡 σ ρᴬ ρᴮ ρᴬᴮ l → 𝓡 σ ρᴬ ρᴮ ρᴬᴮ r → 92 | 𝓡 σ ρᴬ ρᴮ ρᴬᴮ (`ifte b l r) 93 | \end{code} 94 | % 95 | %<*lam> 96 | \begin{code} 97 | lamᴿ : 𝓔ᴿ ρᴬ ρᴮ ρᴬᴮ → ∀ b → 98 | (∀ {Ω} (ρ : Thinning Θ Ω) {vᴮ vᴬᴮ} → rel 𝓥ᴿ σ vᴮ vᴬᴮ → 99 | let σᴬ = th^Env 𝓢ᴬ.th^𝓥 ρᴬ extend ∙ var0ᴬ 100 | σᴮ = th^Env 𝓢ᴮ.th^𝓥 ρᴮ ρ ∙ vᴮ 101 | σᴬᴮ = th^Env 𝓢ᴬᴮ.th^𝓥 ρᴬᴮ ρ ∙ vᴬᴮ 102 | in 𝓡 τ σᴬ σᴮ σᴬᴮ b) → 103 | 𝓡 (σ `→ τ) ρᴬ ρᴮ ρᴬᴮ (`lam b) 104 | \end{code} 105 | % 106 | \begin{code} 107 | private 108 | variable 109 | 𝓢ᴬ : Semantics 𝓥ᴬ 𝓒ᴬ 110 | 𝓢ᴮ : Semantics 𝓥ᴮ 𝓒ᴮ 111 | 𝓢ᴬᴮ : Semantics 𝓥ᴬᴮ 𝓒ᴬᴮ 112 | 𝓔ᴿ : ∀ {Γ Δ Θ} → (Γ ─Env) 𝓥ᴬ Δ → (Δ ─Env) 𝓥ᴮ Θ → (Γ ─Env) 𝓥ᴬᴮ Θ → Set 113 | 𝓥ᴿ : Rel 𝓥ᴮ 𝓥ᴬᴮ 114 | 𝓒ᴿ : Rel 𝓒ᴮ 𝓒ᴬᴮ 115 | module _ (𝓕 : Fusion 𝓢ᴬ 𝓢ᴮ 𝓢ᴬᴮ 𝓔ᴿ 𝓥ᴿ 𝓒ᴿ) where 116 | 117 | open Fusion 𝓕 118 | \end{code} 119 | %<*fundamental:type> 120 | \begin{code} 121 | fusion : 𝓔ᴿ ρᴬ ρᴮ ρᴬᴮ → ∀ t → 𝓡 σ ρᴬ ρᴮ ρᴬᴮ t 122 | \end{code} 123 | % 124 | %<*fundamental:var> 125 | \begin{code} 126 | fusion ρᴿ (`var v) = varᴿ ρᴿ v 127 | \end{code} 128 | % 129 | %<*fundamental:lam> 130 | \begin{code} 131 | fusion ρᴿ (`lam b) = lamᴿ ρᴿ b $ λ ren vᴿ → fusion (th^𝓔ᴿ ρᴿ ren ∙ᴿ vᴿ) b 132 | \end{code} 133 | % 134 | %<*fundamental:base> 135 | \begin{code} 136 | fusion ρᴿ `one = oneᴿ ρᴿ 137 | fusion ρᴿ `tt = ttᴿ ρᴿ 138 | fusion ρᴿ `ff = ffᴿ ρᴿ 139 | \end{code} 140 | % 141 | %<*fundamental:struct> 142 | \begin{code} 143 | fusion ρᴿ (`app f t) = appᴿ ρᴿ f t (fusion ρᴿ f) (fusion ρᴿ t) 144 | fusion ρᴿ (`ifte b l r) = ifteᴿ ρᴿ b l r (fusion ρᴿ b) (fusion ρᴿ l) (fusion ρᴿ r) 145 | \end{code} 146 | % 147 | -------------------------------------------------------------------------------- /doc/type-scope-semantics.agda/Semantics/Printing.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | {-# OPTIONS --safe --sized-types #-} 3 | 4 | module Semantics.Printing where 5 | 6 | open import Size using (∞) 7 | open import Data.List.Base as List using (List; []; _∷_) 8 | open import Data.Var hiding (_<$>_; get) 9 | open import Data.Environment hiding (_<$>_; _++^Env_) 10 | open import Syntax.Type 11 | open import Syntax.Calculus 12 | open import Semantics.Specification 13 | 14 | open import Function 15 | open import Data.Nat 16 | open import Data.Product hiding (map) 17 | open import Data.Char using (Char) 18 | open import Data.String using (String; _++_; concat; fromList; toList) 19 | open import Data.Nat.Show as NatShow 20 | open import Data.List.NonEmpty as List⁺ using (List⁺; _∷_) 21 | 22 | open import Codata.Thunk 23 | open import Codata.Stream as Stream using (Stream ; head ; tail ; zipWith ; _∷_) 24 | open import Category.Monad.State 25 | open RawIMonadState (StateMonadState (Stream String _)) hiding (zipWith ; pure) 26 | open import Relation.Unary 27 | open import Relation.Binary.PropositionalEquality using (_≡_; refl) 28 | 29 | private 30 | variable 31 | I : Set 32 | Γ Δ : List I 33 | α β σ τ : I 34 | A B : Set 35 | 36 | \end{code} 37 | %<*monad> 38 | \begin{code} 39 | Fresh : Set → Set 40 | Fresh = State (Stream String ∞) 41 | \end{code} 42 | % 43 | 44 | %<*valprint> 45 | \begin{code} 46 | record Wrap A (σ : I) (Γ : List I) : Set where 47 | constructor MkW; field getW : A 48 | \end{code} 49 | % 50 | \begin{code} 51 | open Wrap public 52 | 53 | th^Wrap : Thinnable {I} (Wrap A σ) 54 | th^Wrap w ρ = MkW (getW w) 55 | 56 | map^Wrap : (A → B) → ∀[ Wrap A σ ⇒ Wrap B σ ] 57 | map^Wrap f (MkW a) = MkW (f a) 58 | 59 | reindex^Wrap : Wrap A σ Γ → Wrap A τ Δ 60 | reindex^Wrap (MkW w) = MkW w 61 | \end{code} 62 | %<*name> 63 | \begin{code} 64 | Name : I ─Scoped 65 | Name = Wrap String 66 | \end{code} 67 | % 68 | %<*printer> 69 | \begin{code} 70 | Printer : I ─Scoped 71 | Printer = Wrap (Fresh String) 72 | \end{code} 73 | % 74 | \begin{code} 75 | 76 | parens : String → String 77 | parens str = "(" ++ str ++ ")" 78 | 79 | unwords : List String → String 80 | unwords = concat ∘ List.intersperse " " 81 | 82 | \end{code} 83 | %<*fresh> 84 | \begin{code} 85 | fresh : ∀ σ → Fresh (Name σ (σ ∷ Γ)) 86 | fresh σ = do 87 | names ← get 88 | put (tail names) 89 | return (MkW (head names)) 90 | \end{code} 91 | % 92 | 93 | %<*semprint> 94 | \begin{code} 95 | Printing : Semantics Name Printer 96 | Printing = record { th^𝓥 = th^Wrap; var = var; app = app; lam = lam 97 | ; one = one; tt = tt; ff = ff; ifte = ifte } 98 | \end{code} 99 | % 100 | \begin{code} 101 | where 102 | \end{code} 103 | %<*printvar> 104 | \begin{code} 105 | var : ∀[ Name σ ⇒ Printer σ ] 106 | var = map^Wrap return 107 | \end{code} 108 | % 109 | %<*printlam> 110 | \begin{code} 111 | lam : ∀[ □ (Name σ ⇒ Printer τ) ⇒ Printer (σ `→ τ) ] 112 | lam {σ} mb = MkW do 113 | x ← fresh σ 114 | b ← getW (mb extend x) 115 | return ("λ" ++ getW x ++ ". " ++ b) 116 | \end{code} 117 | % 118 | %<*printcons> 119 | \begin{code} 120 | one : ∀[ Printer `Unit ] 121 | one = MkW (return "()") 122 | 123 | tt : ∀[ Printer `Bool ] 124 | tt = MkW (return "true") 125 | 126 | ff : ∀[ Printer `Bool ] 127 | ff = MkW (return "false") 128 | \end{code} 129 | % 130 | %<*printstruct> 131 | \begin{code} 132 | app : ∀[ Printer (σ `→ τ) ⇒ Printer σ ⇒ Printer τ ] 133 | app mf mt = MkW do 134 | f ← getW mf 135 | t ← getW mt 136 | return (f ++ parens t) 137 | 138 | ifte : ∀[ Printer `Bool ⇒ Printer σ ⇒ Printer σ ⇒ Printer σ ] 139 | ifte mb ml mr = MkW do 140 | b ← getW mb 141 | l ← getW ml 142 | r ← getW mr 143 | return (unwords ("if" ∷ parens b ∷ "then" ∷ parens l 144 | ∷ "else" ∷ parens r ∷ []) ) 145 | \end{code} 146 | % 147 | \begin{code} 148 | 149 | alphabetWithSuffix : String → List⁺ String 150 | alphabetWithSuffix suffix = List⁺.map (λ c → fromList (c ∷ []) ++ suffix) 151 | $′ 'a' ∷ toList "bcdefghijklmnopqrstuvwxyz" 152 | 153 | allNats : Stream ℕ _ 154 | allNats = Stream.iterate suc 0 155 | 156 | names : Stream String _ 157 | names = Stream.concat 158 | $′ Stream.map alphabetWithSuffix 159 | $′ "" ∷ λ where .force → Stream.map NatShow.show allNats 160 | 161 | instance _ = rawIApplicative 162 | 163 | \end{code} 164 | %<*printclosed> 165 | \begin{code} 166 | print : ∀ σ → Term σ [] → String 167 | print σ t = proj₁ (getW printer names) where 168 | 169 | printer : Printer σ [] 170 | printer = semantics Printing ε t 171 | \end{code} 172 | % 173 | %<*test> 174 | \begin{code} 175 | _ : print (σ `→ σ) (`lam (`var z)) ≡ "λa. a" 176 | _ = refl 177 | \end{code} 178 | % 179 | -------------------------------------------------------------------------------- /doc/type-scope-semantics.tex/normalization/betaiota.tex: -------------------------------------------------------------------------------- 1 | \section{Normalisation by Evaluation for βι} 2 | 3 | The decision to apply the η-rule lazily as we have done at the beginning of 4 | this chapter can be pushed even further: one may forgo using the ξ-rule too 5 | and simply perform weak-head normalisation. This drives computation only when 6 | absolutely necessary, e.g. when two terms compared for equality have matching 7 | head constructors and one needs to inspect these constructors' arguments to 8 | conclude. 9 | 10 | For that purpose, we introduce an inductive family describing terms in 11 | weak-head normal forms. 12 | 13 | \subsection{Weak-Head Normal Forms} 14 | 15 | A weak-head normal form (respectively a weak-head neutral form) is a term 16 | which has been evaluated just enough to reveal a head constructor 17 | (respectively to reach a stuck elimination). There are no additional 18 | constraints on the subterms: a λ-headed term is in weak-head normal form 19 | no matter the shape of its body. Similarly an application composed of a 20 | variable as the function and a term as the argument is in weak-head neutral 21 | form no matter what the argument looks like. This means in particular 22 | that unlike with \AD{Ne} and \AD{Nf} there is no mutual dependency between 23 | the definitions of \AD{WHNE} (defined first) and \AD{WHNF}. 24 | 25 | \begin{figure}[h] 26 | \ExecuteMetaData[type-scope-semantics.agda/Syntax/WeakHead.tex]{weakhead} 27 | \caption{Weak-Head Normal and Neutral Forms\label{fig:weakhead}} 28 | \end{figure} 29 | 30 | Naturally, it is possible to define the thinnings 31 | \AF{th\textasciicircum{}WHNE} and \AF{th\textasciicircum{}WHNF} 32 | as well as erasure 33 | functions \AF{erase\textasciicircum{}WHNE} and \AF{erase\textasciicircum{}WHNF} 34 | with codomain \AD{Term}. We omit their simple definitions here. 35 | 36 | \subsection{Model Construction} 37 | 38 | The model construction is much like the previous one except 39 | that source terms are now stored in the model too. This means that 40 | from an element of the model, one can pick either the reduced version 41 | of the input term (i.e. a stuck term or the term's computational 42 | content) or the original. We exploit this ability most 43 | notably in reification where once we have obtained either a 44 | head constructor or a head variable, no subterm needs to 45 | be evaluated. 46 | 47 | \begin{figure}[h] 48 | \ExecuteMetaData[type-scope-semantics.agda/Semantics/NormalisationByEvaluation/BetaIota.tex]{model} 49 | \caption{Model Definition for Computing Weak-Head Normal Forms\label{fig:betaiotamodel}} 50 | \end{figure} 51 | 52 | \AF{Thinnable} can be defined rather straightforwardly based on the template 53 | provided in the previous sections: once more all the notions used in the model 54 | definition are \AF{Thinnable} themselves. Reflection and reification also 55 | follow the same recipe as in the previous section. 56 | 57 | The application and conditional branching rules are more 58 | interesting. One important difference with respect to the previous 59 | section is that we do not grow the spine of a stuck term using 60 | reified versions of its arguments but rather the corresponding 61 | \emph{source} term. Thus staying true to the idea that we only head 62 | reduce enough to expose either a constructor or a variable and let 63 | the other subterms untouched. 64 | 65 | \begin{figure}[h] 66 | \ExecuteMetaData[type-scope-semantics.agda/Semantics/NormalisationByEvaluation/BetaIota.tex]{app} 67 | \ExecuteMetaData[type-scope-semantics.agda/Semantics/NormalisationByEvaluation/BetaIota.tex]{ifte} 68 | \caption{Semantical Counterparts of \AIC{`app}~and \AIC{`ifte}\label{fig:betaiotaappifte}} 69 | \end{figure} 70 | 71 | The semantical counterpart of \AIC{`lam} is also slightly trickier than 72 | before. Indeed, we need to recover the source term the value corresponds 73 | to. Luckily we know it has to be λ-headed and once we have introduced a 74 | fresh variable with \AIC{`lam}, we can project out the source term of 75 | the body evaluated using this fresh variable as a placeholder value. 76 | 77 | \begin{figure}[h] 78 | \ExecuteMetaData[type-scope-semantics.agda/Semantics/NormalisationByEvaluation/BetaIota.tex]{lam} 79 | \caption{Semantical Counterparts of \AIC{`lam}\label{fig:betaiotalam}} 80 | \end{figure} 81 | 82 | We can finally put together all of these semantic counterparts to 83 | obtain a \AR{Semantics} corresponding to weak-head normalisation. 84 | We omit the now self-evident definition of \AF{NBE} where \ARF{embed} 85 | is obtained by using \AF{reflect}. 86 | 87 | We can once more run our test and observe that it simply outputs the 88 | term it was fed. Indeed our example is λ-headed, meaning that it is 89 | already in weak-head normal form and that the normaliser does not need 90 | to do any work. 91 | 92 | \begin{figure}[h] 93 | \ExecuteMetaData[type-scope-semantics.agda/Semantics/NormalisationByEvaluation/BetaIota.tex]{test} 94 | \caption{Running example: the βι case}\label{fig:betaiotatest} 95 | \end{figure} 96 | -------------------------------------------------------------------------------- /doc/generic-syntax.agda/Generic/Examples/Elaboration.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | module Generic.Examples.Elaboration where 3 | 4 | import Level 5 | open import Size 6 | open import Function 7 | import Category.Monad as CM 8 | open import Data.Bool 9 | open import Data.Product as Prod 10 | open import Data.List hiding ([_] ; lookup) 11 | open import Data.List.Relation.Unary.All as All hiding (lookup) 12 | open import Data.Maybe as Maybe using (Maybe; nothing; just) 13 | import Data.Maybe.Categorical as MC 14 | 15 | open import Generic.Examples.TypeChecking 16 | using (Mode; LangC; Lang; Type; module M) 17 | ; open Mode; open LangC; open Type; open M 18 | 19 | open import Generic.Examples.STLC 20 | 21 | open import Relation.Unary hiding (_∈_) 22 | open import Data.Var hiding (_<$>_) 23 | open import Data.Environment hiding (_<$>_) 24 | open import Generic.Syntax 25 | open import Generic.Semantics 26 | 27 | open import Relation.Binary.PropositionalEquality hiding ([_]) 28 | 29 | \end{code} 30 | %<*typing> 31 | \begin{code} 32 | Typing : List Mode → Set 33 | Typing = All (const Type) 34 | \end{code} 35 | % 36 | \begin{code} 37 | 38 | private 39 | variable 40 | σ : Type 41 | m : Mode 42 | ms ns : List Mode 43 | \end{code} 44 | %<*fromtyping> 45 | \begin{code} 46 | fromTyping : Typing ms → List Type 47 | fromTyping [] = [] 48 | fromTyping (σ ∷ Γ) = σ ∷ fromTyping Γ 49 | \end{code} 50 | % 51 | %<*elab> 52 | \begin{code} 53 | Elab : Type ─Scoped → Type → (ms : List Mode) → Typing ms → Set 54 | Elab T σ _ Γ = T σ (fromTyping Γ) 55 | \end{code} 56 | % 57 | %<*typemode> 58 | \begin{code} 59 | Type- : Mode ─Scoped 60 | Type- Check ms = ∀ Γ → (σ : Type) → Maybe (Elab (Tm STLC ∞) σ ms Γ) 61 | Type- Infer ms = ∀ Γ → Maybe (Σ[ σ ∈ Type ] Elab (Tm STLC ∞) σ ms Γ) 62 | \end{code} 63 | % 64 | %<*varmode> 65 | \begin{code} 66 | data Var- : Mode ─Scoped where 67 | `var : (infer : ∀ Γ → Σ[ σ ∈ Type ] Elab Var σ ms Γ) → Var- Infer ms 68 | \end{code} 69 | % 70 | \begin{code} 71 | open import Data.List.Relation.Unary.Any hiding (lookup) 72 | open import Data.List.Membership.Propositional 73 | 74 | toVar : m ∈ ms → Var m ms 75 | toVar (here refl) = z 76 | toVar (there v) = s (toVar v) 77 | 78 | fromVar : Var m ms → m ∈ ms 79 | fromVar z = here refl 80 | fromVar (s v) = there (fromVar v) 81 | 82 | coth^Typing : Typing ns → Thinning ms ns → Typing ms 83 | coth^Typing Δ ρ = All.tabulate (λ x∈Γ → All.lookup Δ (fromVar (lookup ρ (toVar x∈Γ)))) 84 | 85 | lookup-fromVar : ∀ Δ (v : Var m ms) → Var (All.lookup Δ (fromVar v)) (fromTyping Δ) 86 | lookup-fromVar (_ ∷ _) z = z 87 | lookup-fromVar (_ ∷ _) (s v) = s (lookup-fromVar _ v) 88 | 89 | erase^coth : ∀ ms Δ (ρ : Thinning ms ns) → 90 | Var σ (fromTyping (coth^Typing Δ ρ)) → Var σ (fromTyping Δ) 91 | erase^coth [] Δ ρ () 92 | erase^coth (m ∷ ms) Δ ρ z = lookup-fromVar Δ (lookup ρ z) 93 | erase^coth (m ∷ ms) Δ ρ (s v) = erase^coth ms Δ (select extend ρ) v 94 | 95 | th^Var- : Thinnable (Var- m) 96 | th^Var- (`var infer) ρ = `var λ Δ → 97 | let (σ , v) = infer (coth^Typing Δ ρ) in 98 | σ , erase^coth _ Δ ρ v 99 | 100 | open Semantics 101 | \end{code} 102 | %<*equal> 103 | \begin{code} 104 | _==_ : (σ τ : Type) → Maybe (σ ≡ τ) 105 | α == α = just refl 106 | (σ₁ `→ τ₁) == (σ₂ `→ τ₂) = do 107 | refl ← σ₁ == σ₂ 108 | refl ← τ₁ == τ₂ 109 | return refl 110 | _ == _ = nothing 111 | \end{code} 112 | % 113 | %<*arrow> 114 | \begin{code} 115 | data Arrow : Type → Set where 116 | _`→_ : (σ τ : Type) → Arrow (σ `→ τ) 117 | 118 | isArrow : ∀ σ → Maybe (Arrow σ) 119 | isArrow α = nothing 120 | isArrow (σ `→ τ) = just (σ `→ τ) 121 | \end{code} 122 | % 123 | %<*app> 124 | \begin{code} 125 | APP : ∀[ Type- Infer ⇒ Type- Check ⇒ Type- Infer ] 126 | APP f t Γ = do 127 | (σ`→τ , F) ← f Γ 128 | (σ `→ τ) ← isArrow σ`→τ 129 | T ← t Γ σ 130 | return (τ , `app F T) 131 | \end{code} 132 | % 133 | %<*lam> 134 | \begin{code} 135 | VAR0 : Var- Infer (Infer ∷ ms) 136 | VAR0 = `var λ where (σ ∷ _) → (σ , z) 137 | 138 | LAM : ∀[ Kripke Var- Type- (Infer ∷ []) Check ⇒ Type- Check ] 139 | LAM b Γ σ`→τ = do 140 | (σ `→ τ) ← isArrow σ`→τ 141 | B ← b (bind Infer) (ε ∙ VAR0) (σ ∷ Γ) τ 142 | return (`lam B) 143 | \end{code} 144 | % 145 | %<*emb> 146 | \begin{code} 147 | EMB : ∀[ Type- Infer ⇒ Type- Check ] 148 | EMB t Γ σ = do 149 | (τ , T) ← t Γ 150 | refl ← σ == τ 151 | return T 152 | \end{code} 153 | % 154 | %<*elaborate> 155 | \begin{code} 156 | Elaborate : Semantics Lang Var- Type- 157 | Elaborate .th^𝓥 = th^Var- 158 | Elaborate .var = λ where (`var infer) Γ → just (map₂ `var (infer Γ)) 159 | Elaborate .alg = λ where 160 | (App , f , t , refl) → APP f t 161 | (Lam , b , refl) → LAM b 162 | (Emb , t , refl) → EMB t 163 | (Cut σ , t , refl) → λ Γ → (σ ,_) <$> t Γ σ 164 | \end{code} 165 | % 166 | \end{code} 167 | -------------------------------------------------------------------------------- /doc/generic-syntax.tex/catalogue/unrolling.tex: -------------------------------------------------------------------------------- 1 | \subsection{Binding as Self-Reference: Representing Cyclic Structures}\label{def:colist} 2 | 3 | Ghani, Hamana, Uustalu and Vene~\citeyear{ghani2006representing} have 4 | demonstrated how Altenkirch and Reus' type-level de Bruijn 5 | indices~\citeyear{altenkirch1999monadic} can be used to represent 6 | potentially cyclic structures by a finite object. In their 7 | representation each bound variable is a pointer to the node 8 | that introduced it. Given that we are, at the top-level, only 9 | interested in structures with no ``dangling pointers'', we introduce 10 | the notation \AF{TM} \AB{d} to mean closed terms (i.e. terms of type 11 | \AD{Tm} \AB{d} \AF{∞} \AIC{[]}). 12 | 13 | A basic example of such a structure is a potentially cyclic list which 14 | offers a choice of two constructors: \AIC{[]} which ends the list and 15 | \AIC{\_::\_} which combines a head and a tail but also acts as a binder 16 | for a self-reference; these pointers can be used by using the \AIC{var} 17 | constructor which we have renamed \AIC{↶} (pronounced ``backpointer'') 18 | to match the domain-specific meaning. 19 | We can see this approach in action in the examples 20 | \AF{[0, 1]} and \AF{01↺} (pronounced ``0-1-cycle'') which describe 21 | respectively a finite list containing 22 | 0 followed by 1 and a cyclic list starting with 0, then 1, and then 23 | repeating the whole list again by referring to the first cons cell 24 | represented here by the de Bruijn variable 1 (i.e. \AIC{s} \AIC{z}). 25 | 26 | \begin{figure}[h] 27 | \begin{minipage}{0.55\textwidth} 28 | \ExecuteMetaData[Generic/Examples/Colist.tex]{clistD} 29 | \ExecuteMetaData[Generic/Examples/Colist.tex]{clistpat} 30 | \end{minipage}\hfill 31 | \begin{minipage}{0.35\textwidth} 32 | \ExecuteMetaData[Generic/Examples/Colist.tex]{zeroones} 33 | \end{minipage} 34 | \caption{Potentially Cyclic Lists: Description, Pattern Synonyms and Examples} 35 | \label{fig:examplecyclic} 36 | \end{figure} 37 | 38 | These finite representations are interesting in their own right 39 | and we can use the generic semantics framework defined earlier 40 | to manipulate them. A basic building block is the \AF{unroll} 41 | function which takes a closed tree, exposes its top node and 42 | unrolls any cycle which has it as its starting point. We can 43 | decompose it using the \AF{plug} function which, given a closed 44 | and an open term, closes the latter by plugging the former at 45 | each free \AIC{`var} leaf. Noticing that \AF{plug}'s fundamental nature 46 | is that of substituting a term for each leaf, it makes sense to 47 | implement it by re-using the \AF{Substitution} semantics we already have. 48 | 49 | \begin{figure}[h] 50 | \begin{minipage}{0.52\textwidth} 51 | \ExecuteMetaData[Generic/Cofinite.tex]{plug} 52 | \end{minipage}\hspace{2em} 53 | \begin{minipage}{0.43\textwidth} 54 | \ExecuteMetaData[Generic/Cofinite.tex]{unroll} 55 | \end{minipage} 56 | \caption{Plug and Unroll: Exposing a Cyclic Tree's Top Layer} 57 | \end{figure} 58 | 59 | However, one thing still out of our reach with our current tools 60 | is the underlying co-finite trees these finite objects are meant 61 | to represent. We start by defining the coinductive type 62 | corresponding to them as the greatest fixpoint of a notion of 63 | layer. One layer of a co-finite tree is precisely given by the 64 | meaning of its description where we completely ignore the binding 65 | structure. We show with \AF{01⋯} the infinite list that 66 | corresponds to the unfolding of the example \AF{01↺} given above 67 | in Figure~\ref{fig:examplecyclic}. 68 | 69 | %%%The definition proceeds by copattern-matching as introduced in 70 | %%%\cite{abel2013copatterns} and showcased in \cite{thibodeau2016case}. 71 | 72 | \begin{figure}[h] 73 | \begin{minipage}{0.5\textwidth} 74 | \ExecuteMetaData[Generic/Cofinite.tex]{cotm} 75 | \end{minipage}\hfill 76 | \begin{minipage}{0.4\textwidth} 77 | \ExecuteMetaData[Generic/Examples/Colist.tex]{zeroones2} 78 | \end{minipage} 79 | \caption{Co-finite Trees: Definition and Example} 80 | \end{figure} 81 | 82 | We can then make the connection between potentially cyclic 83 | structures and the co-finite trees formal by giving an \AF{unfold} 84 | function which, given a closed term, produces its unfolding. 85 | The definition proceeds by unrolling the term's top layer and 86 | co-recursively unfolding all the subterms. 87 | 88 | \begin{figure}[h] 89 | \ExecuteMetaData[Generic/Cofinite.tex]{unfold} 90 | \caption{Generic Unfold of Potentially Cyclic Structures} 91 | \end{figure} 92 | 93 | Even if the 94 | powerful notion of semantics described in Section~\ref{section:semantics} 95 | cannot encompass all the traversals we may be interested in, 96 | it provides us with reusable building blocks: the definition 97 | of \AF{unfold} was made very simple by reusing the generic 98 | program \AF{fmap} and the \AF{Substitution} semantics whilst 99 | the definition of \AR{∞Tm} was made easy by reusing \AF{⟦\_⟧}. 100 | -------------------------------------------------------------------------------- /doc/generic-syntax.tex/catalogue/desugaring.tex: -------------------------------------------------------------------------------- 1 | \section{Sugar and Desugaring as a Semantics}\label{section:letbinding} 2 | 3 | One of the advantages of having a universe of programming language 4 | descriptions is the ability to concisely define an \emph{extension} 5 | of an existing language by using \AD{Desc}ription transformers 6 | grafting extra constructors à la W. Swiestra~(\citeyear{swierstra_2008}). 7 | This is made extremely simple by the disjoint sum combinator 8 | \AF{\_`+\_} which we defined in Figure~\ref{figure:descsum}. 9 | An example of such an extension is the addition of let-bindings to 10 | an existing language. 11 | 12 | Let bindings allow the user to avoid repeating themselves by naming 13 | sub-expressions and then using these names to refer to the associated 14 | terms. Preprocessors adding these types of mechanisms to existing 15 | languages (from C to CSS) are rather popular. In Figure~\ref{defn:letD}, 16 | we introduce a description \AD{Let} which can be used to extend any 17 | language description \AB{d} to a language with let-bindings 18 | (\AB{d} \AF{`+} \AF{Let}). 19 | 20 | \begin{figure}[h] 21 | \begin{minipage}[t]{0.45\textwidth} 22 | \ExecuteMetaData[generic-syntax.agda/Generic/Syntax/LetBinder.tex]{letcode} 23 | \end{minipage} 24 | \begin{minipage}[t]{0.45\textwidth} 25 | \ExecuteMetaData[generic-syntax.agda/Generic/Syntax/LetBinder.tex]{letpattern} 26 | \end{minipage} 27 | \caption{Description of a single let binding, associated pattern synonyms 28 | \label{defn:letD}} 29 | \end{figure} 30 | 31 | This description states that a let-binding node stores a pair of types 32 | \AB{σ} and \AB{τ} and two subterms. First comes the let-bound 33 | expression of type \AB{σ} and second comes the body of the let which 34 | has type \AB{τ} in a context extended with a fresh variable of type 35 | \AB{σ}. This defines a term of type \AB{τ}. 36 | 37 | In a dependently typed language, a type may depend on a value which 38 | in the presence of let bindings may be a variable standing for an 39 | expression. The user naturally does not want it to make any difference 40 | whether they used a variable referring to a let-bound expression or 41 | the expression itself. Various typechecking strategies can accommodate 42 | this expectation: in Coq~(\cite{Coq:manual}) let bindings are primitive 43 | constructs of the language and have their own typing and reduction 44 | rules whereas in Agda they are elaborated away to the core language 45 | by inlining. 46 | 47 | This latter approach to extending a language \AB{d} with let bindings 48 | by inlining them before typechecking can be implemented generically as 49 | a semantics over (\AB{d} \AF{`+} \AF{Let}). For this semantics, values 50 | in the environment and computations are both let-free terms. The algebra 51 | of the semantics can be defined by parts thanks to \AF{case}, the eliminator 52 | for \AF{\_`+\_} defined in Figure~\ref{figure:descsum}: 53 | the old constructors are kept the same by 54 | interpreting them using the generic substitution algebra (\AF{Substitution}); 55 | whilst the let-binder precisely provides the extra value to be added to the 56 | environment. 57 | 58 | \begin{figure}[h] 59 | \ExecuteMetaData[generic-syntax.agda/Generic/Semantics/Elaboration/LetBinder.tex]{unletcode} 60 | \caption{Desugaring as a \AR{Semantics}\label{defn:UnLet}} 61 | \end{figure} 62 | 63 | The process of removing let binders is then kickstarted with the placeholder 64 | environment \AF{id\textasciicircum{}Tm}~=~\AIC{pack}~\AIC{`var} 65 | of type {(\AB{Γ} \AR{─Env}) (\AD{Tm} \AB{d} ∞) \AB{Γ}}. 66 | 67 | \begin{figure}[h] 68 | \ExecuteMetaData[generic-syntax.agda/Generic/Semantics/Elaboration/LetBinder.tex]{unlet} 69 | \caption{Specialising \AF{semantics} with an environment of placeholder values\label{defn:unlet}} 70 | \end{figure} 71 | 72 | In less than 10 lines of code we have defined a generic extension of 73 | syntaxes with binding together with a semantics which corresponds 74 | to an elaborator translating away this new construct. 75 | In \cref{cps-transformation} we had focused on STLC only 76 | and showed that it is similarly possible to implement a Continuation 77 | Passing Style transformation as the composition of two semantics 78 | à la Hatcliff and Danvy~(\citeyear{hatcliff1994generic}). 79 | The first semantics embeds STLC into Moggi's 80 | Meta-Language~(\citeyear{DBLP:journals/iandc/Moggi91}) and thus fixes 81 | an evaluation order. The second one translates Moggi's ML back into 82 | STLC in terms of explicit continuations with a fixed return type. 83 | 84 | We have demonstrated how easily one can define extensions and combine 85 | them on top of a base language without having to reimplement common 86 | traversals for each one of the intermediate representations. Moreover, 87 | it is possible to define \emph{generic} transformations elaborating 88 | these added features in terms of lower-level ones. This suggests that 89 | this setup could be a good candidate to implement generic compilation 90 | passes and could deal with a framework using a wealth of slightly 91 | different intermediate languages à la 92 | Nanopass~(\cite{Keep:2013:NFC:2544174.2500618}). 93 | -------------------------------------------------------------------------------- /doc/generic-syntax.agda/Generic/Identity.lagda: -------------------------------------------------------------------------------- 1 | \begin{code} 2 | {-# OPTIONS --safe --sized-types #-} 3 | 4 | module Generic.Identity where 5 | 6 | open import Size 7 | open import Agda.Builtin.List 8 | open import Data.Product hiding (zip) 9 | open import Data.Sum 10 | open import Data.Var 11 | open import Data.Relation as R 12 | open import Data.Var.Varlike 13 | open import Data.Environment 14 | open import Generic.Syntax 15 | open import Generic.Semantics 16 | open import Generic.Semantics.Syntactic 17 | open import Generic.Relator 18 | open import Generic.Simulation 19 | open import Generic.Simulation.Syntactic 20 | 21 | open import Function 22 | open import Relation.Binary.PropositionalEquality as PEq 23 | open import Relation.Binary.PropositionalEquality.WithK 24 | open ≡-Reasoning 25 | 26 | private 27 | variable 28 | I : Set 29 | d : Desc I 30 | σ : I 31 | Γ : List I 32 | i j : Size 33 | 34 | data _≅_ {d : Desc I} {σ} {Γ} : {s : Size} → Tm d s σ Γ → {t : Size} → Tm d t σ Γ → Set 35 | ⟨_⟩_≅_ : {d : Desc I} (e : Desc I) → ⟦ e ⟧ (Scope (Tm d i)) σ Γ → ⟦ e ⟧ (Scope (Tm d j)) σ Γ → Set 36 | 37 | data _≅_ {d = d} {σ} {Γ} where 38 | `var : {k l : Var σ Γ} → k ≡ l → `var {s = i} k ≅ `var {s = j} l 39 | `con : {b : ⟦ d ⟧ (Scope (Tm d i)) σ Γ} {c : ⟦ d ⟧ (Scope (Tm d j)) σ Γ} → 40 | ⟨ d ⟩ b ≅ c → `con {s = i} b ≅ `con {s = j} c 41 | 42 | ⟨ e ⟩ b ≅ c = ⟦ e ⟧ᴿ (λ _ _ t u → t ≅ u) b c 43 | 44 | ≅⇒≡ : ∀ {t u : Tm d ∞ σ Γ} → t ≅ u → t ≡ u 45 | ⟨_⟩≅⇒≡ : ∀ e {t u : ⟦ e ⟧ (Scope (Tm d ∞)) σ Γ} → ⟨ e ⟩ t ≅ u → t ≡ u 46 | 47 | ≅⇒≡ (`var eq) = cong `var eq 48 | ≅⇒≡ (`con eq) = cong `con (⟨ _ ⟩≅⇒≡ eq) 49 | 50 | ⟨ `σ A d ⟩≅⇒≡ (refl , eq) = cong -,_ (⟨ d _ ⟩≅⇒≡ eq) 51 | ⟨ `X Δ j d ⟩≅⇒≡ (≅-pr , eq) = cong₂ _,_ (≅⇒≡ ≅-pr) (⟨ d ⟩≅⇒≡ eq) 52 | ⟨ `∎ i ⟩≅⇒≡ eq = ≡-irrelevant _ _ 53 | 54 | module RenId {I : Set} {d : Desc I} where 55 | 56 | ren-id : ∀ (t : Tm d i σ Γ) {ρ} → R.All Eqᴿ Γ ρ (base vl^Var) → ren ρ t ≅ t 57 | ren-body-id : ∀ Δ σ (t : Scope (Tm d i) Δ σ Γ) {ρ} → R.All Eqᴿ Γ ρ (base vl^Var) → 58 | reify vl^Var Δ σ (Semantics.body Ren ρ Δ σ t) ≅ t 59 | 60 | ren-id (`var k) ρᴿ = `var (trans (lookupᴿ ρᴿ k) (lookup-base^Var k)) 61 | ren-id (`con t) ρᴿ = `con (subst₂ (⟨ d ⟩_≅_) (sym $ fmap² d (Semantics.body Ren _) (reify vl^Var) _) (fmap-id d _) 62 | $ liftᴿ d (λ Δ i t → ren-body-id Δ i t ρᴿ) _) 63 | 64 | ren-body-id [] σ t = ren-id t 65 | ren-body-id {Γ = Γ} Δ@(_ ∷ _) σ t {ρ} ρᴿ = ren-id t eqᴿ where 66 | 67 | eqᴿ : R.All Eqᴿ _ (lift vl^Var Δ ρ) (base vl^Var) 68 | lookupᴿ eqᴿ k with split Δ k | inspect (split Δ) k 69 | ... | inj₁ k₁ | PEq.[ eq ] = begin 70 | injectˡ Γ (lookup (base vl^Var) k₁) ≡⟨ cong (injectˡ Γ) (lookup-base^Var k₁) ⟩ 71 | injectˡ Γ k₁ ≡⟨ sym (lookup-base^Var k) ⟩ 72 | lookup (base vl^Var) k ∎ 73 | ... | inj₂ k₂ | PEq.[ eq ] = begin 74 | injectʳ Δ (lookup (base vl^Var) (lookup ρ k₂)) ≡⟨ cong (injectʳ Δ) (lookup-base^Var _) ⟩ 75 | injectʳ Δ (lookup ρ k₂) ≡⟨ cong (injectʳ Δ) (lookupᴿ ρᴿ k₂) ⟩ 76 | injectʳ Δ (lookup (base vl^Var) k₂) ≡⟨ cong (injectʳ Δ) (lookup-base^Var k₂) ⟩ 77 | k ≡⟨ sym (lookup-base^Var k) ⟩ 78 | lookup (base vl^Var) k ∎ 79 | 80 | module _ {I : Set} {d : Desc I} where 81 | 82 | ren-id : ∀ {σ Γ} (t : Tm d ∞ σ Γ) → ren (base vl^Var) t ≡ t 83 | ren-id t = ≅⇒≡ (RenId.ren-id t (packᴿ λ _ → refl)) 84 | 85 | ren-id′ : ∀ {σ Γ} (t : Tm d ∞ σ Γ) → ren (pack id) t ≡ t 86 | ren-id′ t = ≅⇒≡ (RenId.ren-id t (packᴿ λ v → sym (lookup-base^Var v))) 87 | 88 | sub-id : ∀ {σ Γ} (t : Tm d ∞ σ Γ) → sub (base vl^Tm) t ≡ t 89 | sub-id t = begin 90 | sub (base vl^Tm) t ≡⟨ sym $ Simulation.sim RenSub base^VarTmᴿ t ⟩ 91 | ren (base vl^Var) t ≡⟨ ren-id t ⟩ 92 | t ∎ 93 | 94 | sub-id′ : ∀ {σ Γ} (t : Tm d ∞ σ Γ) → sub (pack `var) t ≡ t 95 | sub-id′ t = begin 96 | sub (pack `var) t ≡⟨ sym $ Simulation.sim RenSub (packᴿ λ v → refl) t ⟩ 97 | ren (pack id) t ≡⟨ ren-id′ t ⟩ 98 | t ∎ 99 | 100 | lift[]^Tm : ∀ {Γ Δ} (ρ : (Γ ─Env) (Tm d ∞) Δ) → R.All Eqᴿ Γ ρ (lift vl^Tm [] ρ) 101 | lookupᴿ (lift[]^Tm ρ) k = sym (ren-id (lookup ρ k)) 102 | 103 | 104 | th^base₁^Var : ∀ {Γ Δ} (ρ : Thinning {I} Γ Δ) → R.All Eqᴿ Γ (th^Env th^Var (base vl^Var) ρ) ρ 105 | lookupᴿ (th^base₁^Var ρ) k = cong (lookup ρ) (lookup-base^Var k) 106 | 107 | th^base₂^Var : ∀ {Γ Δ} (ρ : Thinning {I} Γ Δ) → R.All Eqᴿ Γ (th^Env th^Var ρ (base vl^Var)) ρ 108 | lookupᴿ (th^base₂^Var ρ) k = `var-inj (ren-id (`var (lookup ρ k))) 109 | 110 | th^base^Tm : ∀ {Γ Δ} (ρ : (Γ ─Env) (Tm d ∞) Δ) → R.All Eqᴿ Γ (th^Env th^Tm ρ (base vl^Var)) ρ 111 | lookupᴿ (th^base^Tm ρ) k = ren-id (lookup ρ k) 112 | 113 | th^base^s∙z : ∀ {σ Γ} → R.All Eqᴿ _ (th^Env th^Tm (base vl^Tm) (pack s) ∙ `var z) 114 | ((σ ∷ Γ ─Env) (Tm d ∞) (σ ∷ Γ) ∋ pack `var) 115 | lookupᴿ th^base^s∙z z = refl 116 | lookupᴿ th^base^s∙z (s k) = cong (ren (pack s)) (lookup-base^Tm k) 117 | \end{code} 118 | -------------------------------------------------------------------------------- /doc/type-scope-semantics.tex/proof-conclusion.tex: -------------------------------------------------------------------------------- 1 | \chapter{Discussion} 2 | 3 | \section{Summary} 4 | 5 | We have demonstrated that we can exploit the shared structure highlighted by the 6 | introduction of \AR{Semantics} to further alleviate the implementer's pain by 7 | tackling the properties of these \AR{Semantics} in a similarly abstract approach. 8 | 9 | We characterised, using a first logical relation, the traversals which were 10 | producing related outputs provided they were fed related inputs. We then provided 11 | useful instances of this schema thus proving that syntactic traversals are 12 | extensional, that renaming is a special case of substitution or even that 13 | normalisation by evaluation produces equal normal forms provided \AF{PER}-related 14 | evaluation environments. 15 | 16 | A more involved second logical relation gave us a general description of fusion 17 | of traversals where we study triples of semantics such that composing the two 18 | first ones would yield an instance of the third one. We then saw that the four 19 | lemmas about the possible interactions of pairs of renamings and/or substitutions 20 | are all instances of this general framework and can be proven sequentially, the 21 | later results relying on the former ones. We then went on to proving the 22 | substitution lemma for Normalisation by Evaluation. 23 | 24 | \section{Related Work} 25 | 26 | Benton, Hur, Kennedy and McBride's joint work~(\citeyear{benton2012strongly}) was not 27 | limited to defining traversals. They proved fusion lemmas describing the interactions 28 | of renaming and substitution using tactics rather than defining a generic proof framework 29 | like we do. They have also proven the evaluation function of their denotational semantics 30 | correct; however they chose to use propositional equality and to assume function 31 | extensionality rather than resorting to the traditional Partial Equivalence Relation 32 | approach we use. 33 | 34 | Through the careful study of the recursion operator associated to each strictly positive 35 | datatype, Malcolm defined proof principles (\citeyear{DBLP:journals/scp/Malcolm90}) which can 36 | be also used as optimisation principles, just like our fusion principles. Other 37 | optimisations such as deforestation (\cite{wadler1990deforestation}) or transformation 38 | to an equivalent but tail-recursive program (\cite{TomeCortinas:2018:AAM:3240719.3241787}) 39 | have seen a generic treatment. 40 | 41 | 42 | \section{Further work} 43 | 44 | We have now fulfilled one of the three goals we highlighted in 45 | \cref{chapter:program-conclusion}. The question of finding more instances of 46 | \AR{Semantics} and of defining a generic notion of \AR{Semantics} for all 47 | syntaxes with binding is still open. Analogous questions for the proof 48 | frameworks arise naturally. 49 | 50 | 51 | \subsection{Other Instances} 52 | 53 | We have only seen a handful of instances of both the \AR{Simulation} lemma and 54 | the \AR{Fusion} one. They already give us important lemmas when studying the 55 | meta-theory of a language. However there are potential opportunities for more 56 | instances to be defined. 57 | 58 | We would like to know whether the idempotence of normalisation by evaluation 59 | can be proven as a corollary of a fusion lemma for evaluation. This would give 60 | us a nice example of a case where the \ARF{reifyᴬ} is not the identity and 61 | actually does some important work. 62 | 63 | Another important question is whether it is always possible to fuse a 64 | preliminary syntactic traversal followed by a semantics \AB{𝓢} into 65 | a single pass of \AB{𝓢}. 66 | 67 | \subsection{Other Proof Frameworks} 68 | 69 | After implementing \AR{Simulation} and \AR{Fusion}, we can wonder whether there 70 | are any other proof schemas we can make formal. 71 | 72 | As we have explained in \cref{sec:simulationrel}, \AR{Simulation} gives the 73 | \emph{relational} interpretation of evaluation. Defining a similar framework 74 | dealing with a single semantics would give us the \emph{predicate} interpretation 75 | of evaluation. This would give a generalisation of the fundamental lemma of logical 76 | predicates which, once specialised to substitution, would be exactly the traditional 77 | definition one would expect. 78 | 79 | Another possible candidate is an \AR{Identity} framework which would, provided that some 80 | constraints hold of the values in the environment, an evaluation is the identity. So far 81 | we have only related pairs of evaluation results but to prove an identity lemma we would 82 | need to relate the evaluation of a term to the original term itself. Although seemingly 83 | devoid of interest, identity lemmas are useful in practice both when proving or when 84 | optimising away useless traversals. 85 | 86 | We actually faced these two challenges when working on the POPLMark Reloaded challenge 87 | (\cite{poplmark2}). We defined the proper generalisation of the fundamental lemma of 88 | logical predicates but could only give ad-hoc identity lemmas for renaming (and thus 89 | substitution because they are in simulation). 90 | 91 | \subsection{Generic Proof Frameworks} 92 | 93 | In the next part, we are going to define a universe of syntaxes with binding and a generic 94 | notion of semantics over these syntaxes. We naturally want to be able to also prove generic 95 | results about these generic traversals. We are going to have to need to generalise the proof 96 | frameworks to make them syntax generic. 97 | --------------------------------------------------------------------------------