├── .gitignore ├── .travis.yml ├── LICENSE ├── Makefile ├── README.md ├── agda └── Syntax.agda ├── notes └── aim25.txt └── src ├── Makefile ├── acmart.cls ├── core-agda.tex └── macros.tex /.gitignore: -------------------------------------------------------------------------------- 1 | *~ 2 | 3 | ## Core latex/pdflatex auxiliary files: 4 | *.aux 5 | *.lof 6 | *.log 7 | *.lot 8 | *.fls 9 | *.out 10 | *.toc 11 | *.fmt 12 | *.fot 13 | *.cb 14 | *.cb2 15 | 16 | ## Intermediate documents: 17 | *.dvi 18 | *-converted-to.* 19 | # these rules might exclude image files for figures etc. 20 | # *.ps 21 | # *.eps 22 | *.pdf 23 | comment.cut 24 | 25 | ## Agda interface files 26 | *.agdai 27 | 28 | ## Bibliography auxiliary files (bibtex/biblatex/biber): 29 | *.bbl 30 | *.bcf 31 | *.blg 32 | *-blx.aux 33 | *-blx.bib 34 | *.brf 35 | *.run.xml 36 | 37 | ## Build tool auxiliary files: 38 | *.fdb_latexmk 39 | *.synctex 40 | *.synctex.gz 41 | *.synctex.gz(busy) 42 | *.pdfsync 43 | 44 | ## Auxiliary and intermediate files from other packages: 45 | # algorithms 46 | *.alg 47 | *.loa 48 | 49 | # achemso 50 | acs-*.bib 51 | 52 | # amsthm 53 | *.thm 54 | 55 | # beamer 56 | *.nav 57 | *.snm 58 | *.vrb 59 | 60 | # cprotect 61 | *.cpt 62 | 63 | # fixme 64 | *.lox 65 | 66 | #(r)(e)ledmac/(r)(e)ledpar 67 | *.end 68 | *.?end 69 | *.[1-9] 70 | *.[1-9][0-9] 71 | *.[1-9][0-9][0-9] 72 | *.[1-9]R 73 | *.[1-9][0-9]R 74 | *.[1-9][0-9][0-9]R 75 | *.eledsec[1-9] 76 | *.eledsec[1-9]R 77 | *.eledsec[1-9][0-9] 78 | *.eledsec[1-9][0-9]R 79 | *.eledsec[1-9][0-9][0-9] 80 | *.eledsec[1-9][0-9][0-9]R 81 | 82 | # glossaries 83 | *.acn 84 | *.acr 85 | *.glg 86 | *.glo 87 | *.gls 88 | *.glsdefs 89 | 90 | # gnuplottex 91 | *-gnuplottex-* 92 | 93 | # hyperref 94 | *.brf 95 | 96 | # knitr 97 | *-concordance.tex 98 | # TODO Comment the next line if you want to keep your tikz graphics files 99 | *.tikz 100 | *-tikzDictionary 101 | 102 | # listings 103 | *.lol 104 | 105 | # makeidx 106 | *.idx 107 | *.ilg 108 | *.ind 109 | *.ist 110 | 111 | # minitoc 112 | *.maf 113 | *.mlf 114 | *.mlt 115 | *.mtc 116 | *.mtc[0-9] 117 | *.mtc[1-9][0-9] 118 | 119 | # minted 120 | _minted* 121 | *.pyg 122 | 123 | # morewrites 124 | *.mw 125 | 126 | # mylatexformat 127 | *.fmt 128 | 129 | # nomencl 130 | *.nlo 131 | 132 | # sagetex 133 | *.sagetex.sage 134 | *.sagetex.py 135 | *.sagetex.scmd 136 | 137 | # sympy 138 | *.sout 139 | *.sympy 140 | sympy-plots-for-*.tex/ 141 | 142 | # pdfcomment 143 | *.upa 144 | *.upb 145 | 146 | # pythontex 147 | *.pytxcode 148 | pythontex-files-*/ 149 | 150 | # thmtools 151 | *.loe 152 | 153 | # TikZ & PGF 154 | *.dpth 155 | *.md5 156 | *.auxlock 157 | 158 | # todonotes 159 | *.tdo 160 | 161 | # xindy 162 | *.xdy 163 | 164 | # xypic precompiled matrices 165 | *.xyc 166 | 167 | # endfloat 168 | *.ttt 169 | *.fff 170 | 171 | # Latexian 172 | TSWLatexianTemp* 173 | 174 | ## Editors: 175 | # WinEdt 176 | *.bak 177 | *.sav 178 | 179 | # Texpad 180 | .texpadtmp 181 | 182 | # Kile 183 | *.backup 184 | 185 | # KBibTeX 186 | *~[0-9]* 187 | -------------------------------------------------------------------------------- /.travis.yml: -------------------------------------------------------------------------------- 1 | dist: trusty 2 | sudo: required 3 | language: generic 4 | 5 | before_install: 6 | - sudo apt-get -q update 7 | - sudo apt-get install -y texlive-latex-extra texlive-math-extra mathpartir 8 | 9 | script: 10 | - make docs 11 | 12 | deploy: 13 | provider: pages 14 | local_dir: ./docs 15 | project_name: "Specification of Agda" 16 | skip_cleanup: true 17 | target_branch: gh-pages 18 | github_token: $GITHUB_TOKEN # Set in travis-ci.org dashboard 19 | on: 20 | branch: master 21 | -------------------------------------------------------------------------------- /LICENSE: -------------------------------------------------------------------------------- 1 | This is free and unencumbered software released into the public domain. 2 | 3 | Anyone is free to copy, modify, publish, use, compile, sell, or 4 | distribute this software, either in source code form or as a compiled 5 | binary, for any purpose, commercial or non-commercial, and by any 6 | means. 7 | 8 | In jurisdictions that recognize copyright laws, the author or authors 9 | of this software dedicate any and all copyright interest in the 10 | software to the public domain. We make this dedication for the benefit 11 | of the public at large and to the detriment of our heirs and 12 | successors. We intend this dedication to be an overt act of 13 | relinquishment in perpetuity of all present and future rights to this 14 | software under copyright law. 15 | 16 | THE SOFTWARE IS PROVIDED "AS IS", WITHOUT WARRANTY OF ANY KIND, 17 | EXPRESS OR IMPLIED, INCLUDING BUT NOT LIMITED TO THE WARRANTIES OF 18 | MERCHANTABILITY, FITNESS FOR A PARTICULAR PURPOSE AND NONINFRINGEMENT. 19 | IN NO EVENT SHALL THE AUTHORS BE LIABLE FOR ANY CLAIM, DAMAGES OR 20 | OTHER LIABILITY, WHETHER IN AN ACTION OF CONTRACT, TORT OR OTHERWISE, 21 | ARISING FROM, OUT OF OR IN CONNECTION WITH THE SOFTWARE OR THE USE OR 22 | OTHER DEALINGS IN THE SOFTWARE. 23 | 24 | For more information, please refer to 25 | -------------------------------------------------------------------------------- /Makefile: -------------------------------------------------------------------------------- 1 | default : 2 | make -C src 3 | 4 | docs : default clean 5 | cp src/*.pdf docs/ 6 | 7 | .PHONY: clean 8 | clean : 9 | mkdir -p docs/ 10 | rm -f docs/*.pdf 11 | 12 | # EOF 13 | -------------------------------------------------------------------------------- /README.md: -------------------------------------------------------------------------------- 1 | # agda-spec [![Build Status](https://travis-ci.org/agda/agda-spec.svg?branch=master)](https://travis-ci.org/agda/agda-spec) 2 | Specification of Agda. 3 | 4 | * [core-agda.pdf](http://agda.github.io/agda-spec/core-agda.pdf) 5 | -------------------------------------------------------------------------------- /agda/Syntax.agda: -------------------------------------------------------------------------------- 1 | -- TODO: use StrictTotalOrder for QName representation 2 | 3 | module Syntax (QName : Set) where 4 | 5 | open import Data.Nat.Base 6 | open import Data.Nat.Properties using (+-suc; +-identityʳ) 7 | open import Data.List.Base hiding (_∷ʳ_) 8 | open import Relation.Binary.PropositionalEquality using (_≡_; refl; cong; sym; trans) 9 | 10 | -- Well-scoped de Bruijn indices 11 | -- n is the length of the context 12 | 13 | data Var : ℕ → Set where 14 | vzero : ∀{n} → Var (suc n) 15 | vsuc : ∀{n} (x : Var n) → Var (suc n) 16 | 17 | -- Example: in context x,y,z (so n = 3) 18 | -- x is vsuc (vsuc vzero) : Var 3 19 | -- y is vsuc vzero : Var 3 20 | -- z is vzero : Var 3 21 | 22 | -- Qualified names 23 | 24 | DRName = QName -- datatype / record name 25 | FuncName = QName -- function name 26 | ProjName = QName -- projection name (overloaded) 27 | ConsName = QName -- datatype / record constructor name (overloaded) 28 | 29 | -- Data or record constructors 30 | 31 | data ConHead : Set where 32 | dataCon : DRName → ConHead 33 | recCon : DRName → (fs : List ProjName) → ConHead 34 | 35 | -- Sorts are Set₀, Set₁, ... 36 | 37 | data Sort : Set where 38 | uni : (ℓ : ℕ) → Sort 39 | 40 | -- In the following definition of well-scoped syntax, 41 | -- n is always the length of the context. 42 | 43 | mutual 44 | 45 | data Term (n : ℕ) : Set where 46 | var : (x : Var n) (es : Elims n) → Term n 47 | def : (f : FuncName) (es : Elims n) → Term n 48 | con : (c : ConHead) (vs : Args n) → Term n -- Fully applied 49 | lam : (v : Term (suc n)) → Term n 50 | -- Types 51 | dat : (d : DRName) (vs : Args n) → Term n -- Fully applied 52 | sort : (s : Sort) → Term n 53 | pi : (u : Term n) (v : Term (suc n)) → Term n 54 | 55 | data Elim (n : ℕ) : Set where 56 | apply : (u : Term n) → Elim n 57 | proj : (π : ProjName) → Elim n 58 | 59 | Elims : (n : ℕ) → Set 60 | Elims n = List (Elim n) 61 | 62 | Args : (n : ℕ) → Set 63 | Args n = List (Term n) 64 | 65 | -- Example: (A : Set) (x : A) → A 66 | -- is represented by 67 | 68 | exTyId : Term 0 69 | exTyId = pi (sort (uni 0)) (pi (var vzero []) (var (vsuc vzero) [])) 70 | 71 | -- Looking up a field in a field-value collection 72 | 73 | -- TODO: Do we want to ensure |fs| = |vs| ? 74 | data LookupField {a} {A : Set a} : (fs : List ProjName) (vs : List A) (f : ProjName) (v : A) → Set where 75 | 76 | here : ∀{f fs v vs} 77 | → LookupField (f ∷ fs) (v ∷ vs) f v 78 | 79 | there : ∀{f f' fs v v' vs} 80 | → LookupField fs vs f v 81 | → LookupField (f' ∷ fs) (v' ∷ vs) f v 82 | 83 | -- Renamings represented as functions 84 | 85 | Renaming : (Γ Δ : ℕ) → Set 86 | Renaming Γ Δ = Var Δ → Var Γ 87 | 88 | weak : ∀{Γ} → Renaming (suc Γ) Γ 89 | weak x = vsuc x 90 | 91 | -- If Γ ⊢ ρ : Δ 92 | -- then Γ,x ⊢ liftRen ρ : Δ,x. 93 | 94 | liftRen : ∀{Γ Δ} → Renaming Γ Δ → Renaming (suc Γ) (suc Δ) 95 | liftRen ρ vzero = vzero 96 | liftRen ρ (vsuc x) = vsuc (ρ x) 97 | 98 | -- We need sized types to show termination of rename. 99 | {-# TERMINATING #-} 100 | 101 | mutual 102 | rename : ∀{Γ Δ} (ρ : Renaming Γ Δ) (t : Term Δ) → Term Γ 103 | rename ρ (var x es) = var (ρ x) (map (renameElim ρ) es) 104 | rename ρ (def f es) = def f (map (renameElim ρ) es) 105 | rename ρ (con c vs) = con c (map (rename ρ) vs) 106 | rename ρ (lam t) = lam (rename (liftRen ρ) t) 107 | rename ρ (dat d vs) = dat d (map (rename ρ) vs) 108 | rename ρ (sort s) = sort s 109 | rename ρ (pi u v) = pi (rename ρ u) (rename (liftRen ρ) v) 110 | 111 | renameElim : ∀{Γ Δ} (ρ : Renaming Γ Δ) (e : Elim Δ) → Elim Γ 112 | renameElim ρ (apply u) = apply (rename ρ u) 113 | renameElim ρ (proj π) = proj π 114 | 115 | -- Substitutions represented as functions 116 | 117 | -- σ : Subst Γ Δ applies to a term living in context Δ 118 | -- and turns it into a term living in context Γ 119 | 120 | Substitution : (Γ Δ : ℕ) → Set 121 | Substitution Γ Δ = Var Δ → Term Γ 122 | 123 | liftSub : ∀{Γ Δ} (σ : Substitution Γ Δ) → Substitution (suc Γ) (suc Δ) 124 | liftSub σ vzero = var vzero [] 125 | liftSub σ (vsuc x) = rename weak (σ x) 126 | 127 | -- Substitute for the last variable (vzero) 128 | 129 | sg : ∀{Γ} (u : Term Γ) → Substitution Γ (suc Γ) 130 | sg {Γ} u vzero = u 131 | sg {Γ} u (vsuc x) = var x [] 132 | 133 | data All₂ {A B : Set} (R : A → B → Set) : List A → List B → Set where 134 | nil : All₂ R [] [] 135 | cons : ∀{x y xs ys} 136 | → R x y 137 | → All₂ R xs ys 138 | → All₂ R (x ∷ xs) (y ∷ ys) 139 | 140 | mutual 141 | 142 | data Apply {Γ} : (t : Term Γ) (es : Elims Γ) (v : Term Γ) → Set where 143 | 144 | empty : ∀{t} 145 | → Apply t [] t 146 | 147 | var : ∀{x es es'} 148 | → Apply (var x es) es' (var x (es ++ es')) 149 | 150 | def : ∀{f es es'} 151 | → Apply (def f es) es' (def f (es ++ es')) 152 | 153 | proj : ∀{ c πs vs π es u v} 154 | → LookupField πs vs π u 155 | → Apply u es v 156 | → Apply (con (recCon c πs) vs) (proj π ∷ es) v 157 | 158 | lam : ∀{t t' u v es} 159 | → SubstTerm (sg u) t t' 160 | → Apply t' es v 161 | → Apply (lam t) (apply u ∷ es) v 162 | 163 | data SubstTerm {Γ Δ} (σ : Substitution Γ Δ) : Term Δ → Term Γ → Set where 164 | 165 | var : ∀{x : Var Δ} {es : Elims Δ} {v : Term Γ} {es' : Elims Γ} 166 | → All₂ (SubstElim σ) es es' 167 | → Apply (σ x) es' v 168 | → SubstTerm σ (var x es) v 169 | 170 | def : ∀{f : FuncName} {es : Elims Δ} {es' : Elims Γ} 171 | → All₂ (SubstElim σ) es es' 172 | → SubstTerm σ (def f es) (def f es') 173 | 174 | con : ∀{c : ConHead} {vs : Args Δ} {vs' : Args Γ} 175 | → All₂ (SubstTerm σ) vs vs' 176 | → SubstTerm σ (con c vs) (con c vs') 177 | 178 | lam : ∀{v : Term (suc Δ)} {v'} 179 | → SubstTerm (liftSub σ) v v' 180 | → SubstTerm σ (lam v) (lam v') 181 | 182 | dat : ∀{d : DRName} {vs : Args Δ} {vs' : Args Γ} 183 | → All₂ (SubstTerm σ) vs vs' 184 | → SubstTerm σ (dat d vs) (dat d vs') 185 | 186 | pi : ∀{U V U' V'} 187 | → SubstTerm σ U U' 188 | → SubstTerm (liftSub σ) V V' 189 | → SubstTerm σ (pi U V) (pi U' V') 190 | 191 | sort : ∀{s : Sort} 192 | → SubstTerm σ (sort s) (sort s) 193 | 194 | data SubstElim {Γ Δ} (σ : Substitution Γ Δ) : Elim Δ → Elim Γ → Set where 195 | 196 | apply : ∀{u u'} 197 | → SubstTerm σ u u' 198 | → SubstElim σ (apply u) (apply u') 199 | 200 | proj : ∀{π} 201 | → SubstElim σ (proj π) (proj π) 202 | 203 | data FunctionApply {Γ} : (T : Term Γ) (us : Args Γ) (U : Term Γ) → Set where 204 | 205 | empty : ∀{T} 206 | → FunctionApply T [] T 207 | 208 | pi : ∀{u us T U V V'} 209 | → SubstTerm (sg u) V V' 210 | → FunctionApply V' us T 211 | → FunctionApply (pi U V) (u ∷ us) T 212 | 213 | -- n is the length of the context 214 | data Context : (n : ℕ) → Set where 215 | [] : Context zero 216 | _∷ʳ_ : ∀{n} → Context n → Term n → Context (suc n) 217 | 218 | -- n is the size of the outer context; m is the size of the telescope 219 | data Telescope (n : ℕ) : (m : ℕ) → Set where 220 | [] : Telescope n zero 221 | _∷_ : ∀{m} → Term n → Telescope (suc n) m → Telescope n (suc m) 222 | 223 | telescopeSize : ∀{n m} → Telescope n m → ℕ 224 | telescopeSize {m = m} _ = m 225 | 226 | suc-move : (m n : ℕ) → m + suc n ≡ suc m + n 227 | suc-move zero n = refl 228 | suc-move (suc m) n = cong suc (+-suc m n) 229 | 230 | -- Add a telescope to a compatible outer context. 231 | addToContext : ∀{n m} → Context n → Telescope n m → Context (n + m) 232 | addToContext {n} Γ [] rewrite +-identityʳ n = Γ 233 | addToContext {n} {.(suc m)} Γ (_∷_ {m = m} x t) rewrite suc-move n m = addToContext (Γ ∷ʳ x) t 234 | 235 | expandTelescope : ∀ {n m} → Telescope n m → Term (n + m) → Term n 236 | expandTelescope {n} [] T rewrite +-identityʳ n = T 237 | expandTelescope {n} {suc m} (U ∷ Δ) T rewrite suc-move n m = 238 | pi U (expandTelescope Δ T) 239 | 240 | data Pattern (n : ℕ) : Set where 241 | pvariable : Var n → Pattern n 242 | pconstructor : ConsName → List (Pattern n) → Pattern n 243 | pinaccesssible : Term n → Pattern n 244 | 245 | data Copattern (n : ℕ) : Set where 246 | capply : Pattern n → Copattern n 247 | cproj : ProjName → Copattern n 248 | 249 | record ConsDeclaration (n : ℕ) : Set where 250 | constructor mkConsDeclaration 251 | field 252 | consName : ConsName 253 | consType : Term n 254 | 255 | record ProjDeclaration (n : ℕ) : Set where 256 | constructor mkProjDeclaration 257 | field 258 | projName : ProjName 259 | projType : Term n 260 | 261 | record DataSignature : Set where 262 | constructor mkDataSignature 263 | field 264 | name : DRName 265 | {numParams} : ℕ 266 | {numIndices} : ℕ 267 | params : Telescope zero numParams 268 | indices : Telescope numParams numIndices 269 | dsort : Sort 270 | 271 | record DataDefinition : Set where 272 | field 273 | name : DRName 274 | {numParams} : ℕ 275 | params : Telescope zero numParams 276 | constructors : List (ConsDeclaration numParams) 277 | 278 | record RecordSignature : Set where 279 | constructor mkRecordSignature 280 | field 281 | name : DRName 282 | {numParams} : ℕ 283 | params : Telescope zero numParams 284 | dsort : Sort 285 | 286 | record RecordDefinition : Set where 287 | field 288 | name : DRName 289 | {numParams} : ℕ 290 | params : Telescope zero numParams 291 | fconstructor : ConsName 292 | fields : List (ProjDeclaration numParams) 293 | 294 | -- TODO: Update spec 295 | record FuncClause : Set where 296 | field 297 | ctxSize : ℕ 298 | ctx : Context ctxSize 299 | name : FuncName 300 | spine : List (Copattern ctxSize) 301 | rhs : Term ctxSize 302 | 303 | -- TODO: How to represent pattern variables? 304 | data Declaration : Set where 305 | typeSignature : FuncName → Term zero → Declaration 306 | functionClause : FuncClause → Declaration 307 | dataSignature : DataSignature → Declaration 308 | dataDefinition : DataDefinition → Declaration 309 | recordSignature : RecordSignature → Declaration 310 | recordDefinition : RecordDefinition → Declaration 311 | 312 | -- TODO: Update signature declarations 313 | data SignatureDeclaration : Set where 314 | dataSig : (D : DataSignature) → List (ConsDeclaration (DataSignature.numParams D)) → 315 | SignatureDeclaration 316 | recordSig : (R : RecordSignature) → ConsDeclaration (RecordSignature.numParams R) → 317 | List (ProjDeclaration (RecordSignature.numParams R)) → SignatureDeclaration 318 | functionSig : FuncName → Term zero → List FuncClause → SignatureDeclaration 319 | 320 | Signature : Set 321 | Signature = List SignatureDeclaration 322 | 323 | data LookupSig : Signature → SignatureDeclaration → Set where 324 | 325 | here : ∀{sd sds} 326 | → LookupSig (sd ∷ sds) sd 327 | 328 | there : ∀{sd sd' sds} 329 | → LookupSig sds sd 330 | → LookupSig (sd' ∷ sds) sd 331 | 332 | data LookupCons : ∀{m} → List (ConsDeclaration m) → ConsName → Term m → Set where 333 | 334 | here : ∀{c m} {T : Term m} {cs : List (ConsDeclaration m)} 335 | → LookupCons {m} (mkConsDeclaration c T ∷ cs) c T 336 | 337 | there : ∀{c c' m} {T : Term m} {T' : Term m} {cs : List (ConsDeclaration m)} 338 | → LookupCons cs c T 339 | → LookupCons (mkConsDeclaration c' T' ∷ cs) c T 340 | 341 | data LookupProj : ∀{m} → List (ProjDeclaration m) → ProjName → Term m → Set where 342 | 343 | here : ∀{π m} {T : Term m} {πs : List (ProjDeclaration m)} 344 | → LookupProj (mkProjDeclaration π T ∷ πs) π T 345 | 346 | there : ∀ {π π' m} {πs : List (ProjDeclaration m)} {T : Term m} {T' : Term m} 347 | → LookupProj πs π T 348 | → LookupProj (mkProjDeclaration π' T' ∷ πs) π T 349 | 350 | data SigLookup : ∀{m} → Signature → QName → Term m → Set where 351 | func : ∀{Σ f T cls} → LookupSig Σ (functionSig f T cls) → SigLookup Σ f T 352 | dat : ∀{Σ D m m' s cs} {Δ : Telescope zero m} {Δ' : Telescope m m'} → 353 | LookupSig Σ (dataSig (mkDataSignature D {m} {m'} Δ Δ' s) cs) → 354 | SigLookup Σ D (expandTelescope Δ (expandTelescope Δ' (sort s))) 355 | con : ∀{Σ D m m' s cs c T} {Δ : Telescope zero m} {Δ' : Telescope m m'} → 356 | LookupSig Σ (dataSig (mkDataSignature D {m} {m'} Δ Δ' s) cs) → 357 | LookupCons cs c T → 358 | SigLookup Σ c (expandTelescope Δ T) 359 | rec : ∀{Σ R m s c ps} {Δ : Telescope zero m} → 360 | LookupSig Σ (recordSig (mkRecordSignature R Δ s) c ps) → 361 | SigLookup Σ R (expandTelescope Δ (sort s)) 362 | rcon : ∀{Σ R m s c T πs} {Δ : Telescope zero m} → 363 | LookupSig Σ (recordSig (mkRecordSignature R Δ s) (mkConsDeclaration c T) πs) → 364 | SigLookup Σ c (expandTelescope Δ T) 365 | rproj : ∀{Σ R m s c π T πs} {Δ : Telescope zero m} → 366 | LookupSig Σ (recordSig (mkRecordSignature R Δ s) (mkConsDeclaration c T) πs) → 367 | LookupProj πs π T → 368 | SigLookup Σ π (expandTelescope Δ T) 369 | -------------------------------------------------------------------------------- /notes/aim25.txt: -------------------------------------------------------------------------------- 1 | 2017-05-09 2 | 3 | What form should the specification have? 4 | 5 | In the style of "A simple type-theoretic language Mini-TT" 6 | 7 | Core syntax. 8 | 9 | - lambda-Pi 10 | - Universes 11 | - datatypes and functions by pattern matching 12 | - records types, copatterns 13 | 14 | Evaluation rules. 15 | 16 | Typing rules. 17 | 18 | Coverage rules. 19 | 20 | Positivity checking 21 | 22 | Termination checking. 23 | 24 | 25 | 26 | 27 | Extensions 28 | 29 | - universe polymorphism, Setω 30 | 31 | - literals 32 | 33 | - irrelevance 34 | 35 | - abstract 36 | 37 | 38 | 39 | 40 | 41 | 42 | 43 | 44 | TRASH 45 | 46 | --> read off from CheckInternal 47 | 48 | --> datatypes, record types, pattern matching, 49 | 50 | --> 51 | -------------------------------------------------------------------------------- /src/Makefile: -------------------------------------------------------------------------------- 1 | default : core-agda.pdf 2 | 3 | %.pdf : %.tex 4 | pdflatex -interaction nonstopmode -halt-on-error -file-line-error $< 5 | 6 | # EOF 7 | -------------------------------------------------------------------------------- /src/acmart.cls: -------------------------------------------------------------------------------- 1 | %% 2 | %% This is file `acmart.cls', 3 | %% generated with the docstrip utility. 4 | %% 5 | %% The original source files were: 6 | %% 7 | %% acmart.dtx (with options: `class') 8 | %% 9 | %% IMPORTANT NOTICE: 10 | %% 11 | %% For the copyright see the source file. 12 | %% 13 | %% Any modified versions of this file must be renamed 14 | %% with new filenames distinct from acmart.cls. 15 | %% 16 | %% For distribution of the original source see the terms 17 | %% for copying and modification in the file acmart.dtx. 18 | %% 19 | %% This generated file may be distributed as long as the 20 | %% original source files, as listed above, are part of the 21 | %% same distribution. (The sources need not necessarily be 22 | %% in the same archive or directory.) 23 | %% \CharacterTable 24 | %% {Upper-case \A\B\C\D\E\F\G\H\I\J\K\L\M\N\O\P\Q\R\S\T\U\V\W\X\Y\Z 25 | %% Lower-case \a\b\c\d\e\f\g\h\i\j\k\l\m\n\o\p\q\r\s\t\u\v\w\x\y\z 26 | %% Digits \0\1\2\3\4\5\6\7\8\9 27 | %% Exclamation \! Double quote \" Hash (number) \# 28 | %% Dollar \$ Percent \% Ampersand \& 29 | %% Acute accent \' Left paren \( Right paren \) 30 | %% Asterisk \* Plus \+ Comma \, 31 | %% Minus \- Point \. Solidus \/ 32 | %% Colon \: Semicolon \; Less than \< 33 | %% Equals \= Greater than \> Question mark \? 34 | %% Commercial at \@ Left bracket \[ Backslash \\ 35 | %% Right bracket \] Circumflex \^ Underscore \_ 36 | %% Grave accent \` Left brace \{ Vertical bar \| 37 | %% Right brace \} Tilde \~} 38 | 39 | \NeedsTeXFormat{LaTeX2e} 40 | \ProvidesClass{acmart} 41 | [2016/12/03 v1.25 Typesetting articles for Association of 42 | Computing Machinery] 43 | \def\@classname{acmart} 44 | \RequirePackage{xkeyval} 45 | \define@choicekey*+{acmart.cls}{format}[\ACM@format\ACM@format@nr]{% 46 | manuscript, acmsmall, acmlarge, acmtog, sigconf, siggraph, 47 | sigplan, sigchi, sigchi-a}[manuscript]{}{% 48 | \ClassError{\@classname}{The option format must be manuscript, 49 | acmsmall, acmlarge, acmtog, sigconf, siggraph, 50 | sigplan, sigchi or sigchi-a}} 51 | \def\@DeclareACMFormat#1{\DeclareOptionX{#1}{\setkeys{acmart.cls}{format=#1}}} 52 | \@DeclareACMFormat{manuscript} 53 | \@DeclareACMFormat{acmsmall} 54 | \@DeclareACMFormat{acmlarge} 55 | \@DeclareACMFormat{acmtog} 56 | \@DeclareACMFormat{sigconf} 57 | \@DeclareACMFormat{siggraph} 58 | \@DeclareACMFormat{sigplan} 59 | \@DeclareACMFormat{sigchi} 60 | \@DeclareACMFormat{sigchi-a} 61 | \ExecuteOptionsX{format} 62 | \define@boolkey+{acmart.cls}[@ACM@]{screen}[true]{% 63 | \if@ACM@screen 64 | \PackageInfo{\@classname}{Using screen mode}% 65 | \else 66 | \PackageInfo{\@classname}{Not using screen mode}% 67 | \fi}{\PackageError{\@classname}{Option screen can be either true or 68 | false}} 69 | \ExecuteOptionsX{screen=false} 70 | \define@boolkey+{acmart.cls}[@ACM@]{review}[true]{% 71 | \if@ACM@review 72 | \PackageInfo{\@classname}{Using review mode}% 73 | \else 74 | \PackageInfo{\@classname}{Not using review mode}% 75 | \fi}{\PackageError{\@classname}{Option review can be either true or 76 | false}} 77 | \ExecuteOptionsX{review=false} 78 | \define@boolkey+{acmart.cls}[@ACM@]{authorversion}[true]{% 79 | \if@ACM@authorversion 80 | \PackageInfo{\@classname}{Using authorversion mode}% 81 | \else 82 | \PackageInfo{\@classname}{Not using authorversion mode}% 83 | \fi}{\PackageError{\@classname}{Option authorversion can be either true or 84 | false}} 85 | \ExecuteOptionsX{authorversion=false} 86 | \newif\if@ACM@natbib@override 87 | \@ACM@natbib@overridefalse 88 | \define@boolkey+{acmart.cls}[@ACM@]{natbib}[true]{% 89 | \@ACM@natbib@overridetrue 90 | \if@ACM@natbib 91 | \PackageInfo{\@classname}{Explicitly selecting natbib mode}% 92 | \else 93 | \PackageInfo{\@classname}{Explicitly deselecting natbib mode}% 94 | \fi}{\PackageError{\@classname}{Option natbib can be either true or 95 | false}} 96 | \define@boolkey+{acmart.cls}[@ACM@]{anonymous}[true]{% 97 | \if@ACM@anonymous 98 | \PackageInfo{\@classname}{Using anonymous mode}% 99 | \else 100 | \PackageInfo{\@classname}{Not using anonymous mode}% 101 | \fi}{\PackageError{\@classname}{Option anonymous can be either true or 102 | false}} 103 | \ExecuteOptionsX{anonymous=false} 104 | \def\ACM@fontsize{} 105 | \DeclareOptionX{9pt}{\edef\ACM@fontsize{\CurrentOption}} 106 | \DeclareOptionX{10pt}{\edef\ACM@fontsize{\CurrentOption}} 107 | \DeclareOptionX{11pt}{\edef\ACM@fontsize{\CurrentOption}} 108 | \DeclareOptionX{12pt}{\edef\ACM@fontsize{\CurrentOption}} 109 | \DeclareOptionX{draft}{\PassOptionsToClass{\CurrentOption}{amsart}} 110 | \DeclareOptionX{*}{\PassOptionsToClass{\CurrentOption}{amsart}} 111 | \ProcessOptionsX 112 | \ClassInfo{\@classname}{Using format \ACM@format, number \ACM@format@nr} 113 | \newif\if@ACM@manuscript 114 | \newif\if@ACM@journal 115 | \newif\if@ACM@sigchiamode 116 | \ifnum\ACM@format@nr=0\relax 117 | \@ACM@manuscripttrue 118 | \else 119 | \@ACM@manuscriptfalse 120 | \fi 121 | \@ACM@sigchiamodefalse 122 | \ifcase\ACM@format@nr 123 | \relax % manuscript 124 | \@ACM@journaltrue 125 | \or % acmsmall 126 | \@ACM@journaltrue 127 | \or % acmlarge 128 | \@ACM@journaltrue 129 | \or % acmtog 130 | \@ACM@journaltrue 131 | \or % sigconf 132 | \@ACM@journalfalse 133 | \or % siggraph 134 | \@ACM@journalfalse 135 | \or % sigplan 136 | \@ACM@journalfalse 137 | \or % sigchi 138 | \@ACM@journalfalse 139 | \or % sigchi-a 140 | \@ACM@journalfalse 141 | \@ACM@sigchiamodetrue 142 | \fi 143 | \if@ACM@natbib@override\else 144 | \@ACM@natbibtrue 145 | \fi 146 | \ifx\ACM@fontsize\@empty 147 | \ifcase\ACM@format@nr 148 | \relax % manuscript 149 | \def\ACM@fontsize{9pt}% 150 | \or % acmsmall 151 | \def\ACM@fontsize{10pt}% 152 | \or % acmlarge 153 | \def\ACM@fontsize{10pt}% 154 | \or % acmtog 155 | \def\ACM@fontsize{9pt}% 156 | \or % sigconf 157 | \def\ACM@fontsize{9pt}% 158 | \or % siggraph 159 | \def\ACM@fontsize{9pt}% 160 | \or % sigplan 161 | \def\ACM@fontsize{9pt}% 162 | \or % sigchi 163 | \def\ACM@fontsize{10pt}% 164 | \or % sigchi-a 165 | \def\ACM@fontsize{10pt}% 166 | \fi 167 | \fi 168 | \ClassInfo{\@classname}{Using fontsize \ACM@fontsize} 169 | \LoadClass[\ACM@fontsize, reqno]{amsart} 170 | \RequirePackage{microtype} 171 | \RequirePackage{totpages} 172 | \RequirePackage{environ} 173 | \if@ACM@manuscript 174 | \RequirePackage{setspace} 175 | \onehalfspacing 176 | \fi 177 | \if@ACM@natbib 178 | \RequirePackage{natbib} 179 | \renewcommand{\bibsection}{% 180 | \section*{\refname}% 181 | \phantomsection\addcontentsline{toc}{section}{\refname}% 182 | } 183 | \renewcommand{\bibfont}{\bibliofont} 184 | \renewcommand\setcitestyle[1]{ 185 | \@for\@tempa:=#1\do 186 | {\def\@tempb{round}\ifx\@tempa\@tempb 187 | \renewcommand\NAT@open{(}\renewcommand\NAT@close{)}\fi 188 | \def\@tempb{square}\ifx\@tempa\@tempb 189 | \renewcommand\NAT@open{[}\renewcommand\NAT@close{]}\fi 190 | \def\@tempb{angle}\ifx\@tempa\@tempb 191 | \renewcommand\NAT@open{$<$}\renewcommand\NAT@close{$>$}\fi 192 | \def\@tempb{curly}\ifx\@tempa\@tempb 193 | \renewcommand\NAT@open{\{}\renewcommand\NAT@close{\}}\fi 194 | \def\@tempb{semicolon}\ifx\@tempa\@tempb 195 | \renewcommand\NAT@sep{;}\fi 196 | \def\@tempb{colon}\ifx\@tempa\@tempb 197 | \renewcommand\NAT@sep{;}\fi 198 | \def\@tempb{comma}\ifx\@tempa\@tempb 199 | \renewcommand\NAT@sep{,}\fi 200 | \def\@tempb{authoryear}\ifx\@tempa\@tempb 201 | \NAT@numbersfalse\fi 202 | \def\@tempb{numbers}\ifx\@tempa\@tempb 203 | \NAT@numberstrue\NAT@superfalse\fi 204 | \def\@tempb{super}\ifx\@tempa\@tempb 205 | \NAT@numberstrue\NAT@supertrue\fi 206 | \def\@tempb{nobibstyle}\ifx\@tempa\@tempb 207 | \let\bibstyle=\@gobble\fi 208 | \def\@tempb{bibstyle}\ifx\@tempa\@tempb 209 | \let\bibstyle=\@citestyle\fi 210 | \def\@tempb{sort}\ifx\@tempa\@tempb 211 | \def\NAT@sort{\@ne}\fi 212 | \def\@tempb{nosort}\ifx\@tempa\@tempb 213 | \def\NAT@sort{\z@}\fi 214 | \def\@tempb{compress}\ifx\@tempa\@tempb 215 | \def\NAT@cmprs{\@ne}\fi 216 | \def\@tempb{nocompress}\ifx\@tempa\@tempb 217 | \def\NAT@cmprs{\@z}\fi 218 | \def\@tempb{sort&compress}\ifx\@tempa\@tempb 219 | \def\NAT@sort{\@ne}\def\NAT@cmprs{\@ne}\fi 220 | \def\@tempb{mcite}\ifx\@tempa\@tempb 221 | \let\NAT@merge\@ne\fi 222 | \def\@tempb{merge}\ifx\@tempa\@tempb 223 | \@ifnum{\NAT@merge<\tw@}{\let\NAT@merge\tw@}{}\fi 224 | \def\@tempb{elide}\ifx\@tempa\@tempb 225 | \@ifnum{\NAT@merge<\thr@@}{\let\NAT@merge\thr@@}{}\fi 226 | \def\@tempb{longnamesfirst}\ifx\@tempa\@tempb 227 | \NAT@longnamestrue\fi 228 | \def\@tempb{nonamebreak}\ifx\@tempa\@tempb 229 | \def\NAT@nmfmt#1{\mbox{\NAT@up#1}}\fi 230 | \expandafter\NAT@find@eq\@tempa=\relax\@nil 231 | \if\@tempc\relax\else 232 | \expandafter\NAT@rem@eq\@tempc 233 | \def\@tempb{open}\ifx\@tempa\@tempb 234 | \xdef\NAT@open{\@tempc}\fi 235 | \def\@tempb{close}\ifx\@tempa\@tempb 236 | \xdef\NAT@close{\@tempc}\fi 237 | \def\@tempb{aysep}\ifx\@tempa\@tempb 238 | \xdef\NAT@aysep{\@tempc}\fi 239 | \def\@tempb{yysep}\ifx\@tempa\@tempb 240 | \xdef\NAT@yrsep{\@tempc}\fi 241 | \def\@tempb{notesep}\ifx\@tempa\@tempb 242 | \xdef\NAT@cmt{\@tempc}\fi 243 | \def\@tempb{citesep}\ifx\@tempa\@tempb 244 | \xdef\NAT@sep{\@tempc}\fi 245 | \fi 246 | }% 247 | \NAT@@setcites 248 | } 249 | \renewcommand\citestyle[1]{% 250 | \ifcsname bibstyle@#1\endcsname% 251 | \csname bibstyle@#1\endcsname\let\bibstyle\@gobble% 252 | \else% 253 | \@latex@error{Undefined `#1' citestyle}% 254 | \fi 255 | }% 256 | \fi 257 | \newcommand{\bibstyle@acmauthoryear}{% 258 | \setcitestyle{% 259 | authoryear,% 260 | open={(},close={)},citesep={;},% 261 | aysep={},yysep={,},% 262 | notesep={, }}} 263 | \newcommand{\bibstyle@acmnumeric}{% 264 | \setcitestyle{% 265 | numbers,sort&compress,% 266 | open={[},close={]},citesep={,},% 267 | notesep={, }}} 268 | \citestyle{acmnumeric} 269 | \def\@startsection#1#2#3#4#5#6{% 270 | \if@noskipsec \leavevmode \fi 271 | \par 272 | \@tempskipa #4\relax 273 | \@afterindenttrue 274 | \ifdim \@tempskipa <\z@ 275 | \@tempskipa -\@tempskipa \@afterindentfalse 276 | \fi 277 | \if@nobreak 278 | \everypar{}% 279 | \else 280 | \addpenalty\@secpenalty\addvspace\@tempskipa 281 | \fi 282 | \@ifstar 283 | {\@ssect{#3}{#4}{#5}{#6}}% 284 | {\@dblarg{\@sect{#1}{#2}{#3}{#4}{#5}{#6}}}} 285 | \def\@sect#1#2#3#4#5#6[#7]#8{% 286 | \ifnum #2>\c@secnumdepth 287 | \let\@svsec\@empty 288 | \else 289 | \refstepcounter{#1}% 290 | \protected@edef\@svsec{\@seccntformat{#1}\relax}% 291 | \fi 292 | \@tempskipa #5\relax 293 | \ifdim \@tempskipa>\z@ 294 | \begingroup 295 | #6{% 296 | \@hangfrom{\hskip #3\relax\@svsec}% 297 | \interlinepenalty \@M #8\@@par}% 298 | \endgroup 299 | \csname #1mark\endcsname{#7}% 300 | \addcontentsline{toc}{#1}{% 301 | \ifnum #2>\c@secnumdepth \else 302 | \protect\numberline{\csname the#1\endcsname}% 303 | \fi 304 | #7}% 305 | \else 306 | \def\@svsechd{% 307 | #6{\hskip #3\relax 308 | \@svsec #8}% 309 | \csname #1mark\endcsname{#7}% 310 | \addcontentsline{toc}{#1}{% 311 | \ifnum #2>\c@secnumdepth \else 312 | \protect\numberline{\csname the#1\endcsname}% 313 | \fi 314 | #7}}% 315 | \fi 316 | \@xsect{#5}} 317 | \def\@xsect#1{% 318 | \@tempskipa #1\relax 319 | \ifdim \@tempskipa>\z@ 320 | \par \nobreak 321 | \vskip \@tempskipa 322 | \@afterheading 323 | \else 324 | \@nobreakfalse 325 | \global\@noskipsectrue 326 | \everypar{% 327 | \if@noskipsec 328 | \global\@noskipsecfalse 329 | {\setbox\z@\lastbox}% 330 | \clubpenalty\@M 331 | \begingroup \@svsechd \endgroup 332 | \unskip 333 | \@tempskipa #1\relax 334 | \hskip -\@tempskipa 335 | \else 336 | \clubpenalty \@clubpenalty 337 | \everypar{}% 338 | \fi}% 339 | \fi 340 | \ignorespaces} 341 | \def\@seccntformat#1{\csname the#1\endcsname\quad} 342 | \def\@ssect#1#2#3#4#5{% 343 | \@tempskipa #3\relax 344 | \ifdim \@tempskipa>\z@ 345 | \begingroup 346 | #4{% 347 | \@hangfrom{\hskip #1}% 348 | \interlinepenalty \@M #5\@@par}% 349 | \endgroup 350 | \else 351 | \def\@svsechd{#4{\hskip #1\relax #5}}% 352 | \fi 353 | \@xsect{#3}} 354 | \let\@footnotemark@nolink\@footnotemark 355 | \let\@footnotetext@nolink\@footnotetext 356 | \RequirePackage[bookmarksnumbered]{hyperref} 357 | \urlstyle{rm} 358 | \ifcase\ACM@format@nr 359 | \relax % manuscript 360 | \or % acmsmall 361 | \or % acmlarge 362 | \or % acmtog 363 | \or % sigconf 364 | \or % siggraph 365 | \or % sigplan 366 | \urlstyle{sf} 367 | \or % sigchi 368 | \or % sigchi-a 369 | \urlstyle{sf} 370 | \fi 371 | \if@ACM@screen 372 | \hypersetup{colorlinks, 373 | linkcolor=ACMRed, 374 | citecolor=ACMPurple, 375 | urlcolor=ACMDarkBlue, 376 | filecolor=ACMDarkBlue} 377 | \else 378 | \hypersetup{hidelinks} 379 | \fi 380 | \if@ACM@natbib 381 | \let\citeN\cite 382 | \let\cite\citep 383 | \let\citeANP\citeauthor 384 | \let\citeNN\citeyearpar 385 | \let\citeyearNP\citeyear 386 | \let\citeyear\citeyearpar 387 | \let\citeNP\citealt 388 | \def\shortcite#1{\citeyear{#1}} 389 | \DeclareRobustCommand\citeA 390 | {\begingroup\NAT@swafalse 391 | \let\NAT@ctype\@ne\NAT@partrue\NAT@fullfalse\NAT@open\NAT@citetp}% 392 | \providecommand\newblock{}% 393 | \else 394 | \providecommand\citename[1]{#1} 395 | \fi 396 | \def\bibliographystyle#1{% 397 | \ifx\@begindocumenthook\@undefined\else 398 | \expandafter\AtBeginDocument 399 | \fi 400 | {\if@filesw 401 | \immediate\write\@auxout{\string\bibstyle{#1}}% 402 | \fi}} 403 | \RequirePackage{graphicx, xcolor} 404 | \definecolor[named]{ACMBlue}{cmyk}{1,0.1,0,0.1} 405 | \definecolor[named]{ACMYellow}{cmyk}{0,0.16,1,0} 406 | \definecolor[named]{ACMOrange}{cmyk}{0,0.42,1,0.01} 407 | \definecolor[named]{ACMRed}{cmyk}{0,0.90,0.86,0} 408 | \definecolor[named]{ACMLightBlue}{cmyk}{0.49,0.01,0,0} 409 | \definecolor[named]{ACMGreen}{cmyk}{0.20,0,1,0.19} 410 | \definecolor[named]{ACMPurple}{cmyk}{0.55,1,0,0.15} 411 | \definecolor[named]{ACMDarkBlue}{cmyk}{1,0.58,0,0.21} 412 | \RequirePackage{geometry} 413 | \ifcase\ACM@format@nr 414 | \relax % manuscript 415 | \geometry{letterpaper,head=1pc}% 416 | \or % acmsmall 417 | \geometry{twoside=true, 418 | includeheadfoot, head=1pc, foot=2pc, 419 | paperwidth=6.75in, paperheight=10in, 420 | top=58pt, bottom=44pt, inner=46pt, outer=46pt 421 | }% 422 | \or % acmlarge 423 | \geometry{twoside=true, head=1pc, foot=2pc, 424 | paperwidth=8.5in, paperheight=11in, 425 | includeheadfoot, 426 | top=78pt, bottom=114pt, inner=81pt, outer=81pt 427 | }% 428 | \or % acmtog 429 | \geometry{twoside=true, head=1pc, foot=2pc, 430 | paperwidth=8.5in, paperheight=11in, 431 | includeheadfoot, columnsep=24pt, 432 | top=52pt, bottom=75pt, inner=52pt, outer=52pt 433 | }% 434 | \or % sigconf 435 | \geometry{twoside=true, head=1pc, 436 | paperwidth=8.5in, paperheight=11in, 437 | includeheadfoot, columnsep=2pc, 438 | top=57pt, bottom=73pt, inner=54pt, outer=54pt 439 | }% 440 | \or % siggraph 441 | \geometry{twoside=true, head=1pc, 442 | paperwidth=8.5in, paperheight=11in, 443 | includeheadfoot, columnsep=2pc, 444 | top=57pt, bottom=73pt, inner=54pt, outer=54pt 445 | }% 446 | \or % sigplan 447 | \geometry{twoside=true, head=1pc, 448 | paperwidth=8.5in, paperheight=11in, 449 | includeheadfoot=false, columnsep=2pc, 450 | top=1in, bottom=1in, inner=0.75in, outer=0.75in 451 | }% 452 | \or % sigchi 453 | \geometry{twoside=true, head=1pc, 454 | paperwidth=8.5in, paperheight=11in, 455 | includeheadfoot, columnsep=2pc, 456 | top=66pt, bottom=73pt, inner=54pt, outer=54pt 457 | }% 458 | \or % sigchi-a 459 | \geometry{twoside=false, head=1pc, 460 | paperwidth=11in, paperheight=8.5in, 461 | includeheadfoot, marginparsep=72pt, 462 | marginparwidth=170pt, columnsep=20pt, 463 | top=72pt, bottom=72pt, left=314pt, right=72pt 464 | }% 465 | \@mparswitchfalse 466 | \reversemarginpar 467 | \fi 468 | \setlength\parindent{10\p@} 469 | \setlength\parskip{\z@} 470 | \ifcase\ACM@format@nr 471 | \relax % manuscript 472 | \or % acmsmall 473 | \or % acmlarge 474 | \or % acmtog 475 | \setlength\parindent{9\p@}% 476 | \or % sigconf 477 | \or % siggraph 478 | \or % sigplan 479 | \or % sigchi 480 | \or % sigchi-a 481 | \fi 482 | \def\copyrightpermissionfootnoterule{\kern-3\p@ 483 | \hrule \@width \columnwidth \kern 2.6\p@} 484 | \RequirePackage{manyfoot} 485 | \SelectFootnoteRule[2]{copyrightpermission} 486 | \DeclareNewFootnote{copyrightpermission} 487 | \def\footnoterule{\kern-3\p@ 488 | \hrule \@width 4pc \kern 2.6\p@} 489 | \def\endminipage{% 490 | \par 491 | \unskip 492 | \ifvoid\@mpfootins\else 493 | \vskip\skip\@mpfootins 494 | \normalcolor 495 | \unvbox\@mpfootins 496 | \fi 497 | \@minipagefalse %% added 24 May 89 498 | \color@endgroup 499 | \egroup 500 | \expandafter\@iiiparbox\@mpargs{\unvbox\@tempboxa}} 501 | \def\@makefntext{\noindent\@makefnmark} 502 | \if@ACM@sigchiamode 503 | \long\def\@footnotetext#1{\marginpar{% 504 | \reset@font\small 505 | \interlinepenalty\interfootnotelinepenalty 506 | \protected@edef\@currentlabel{% 507 | \csname p@footnote\endcsname\@thefnmark 508 | }% 509 | \color@begingroup 510 | \@makefntext{% 511 | \rule\z@\footnotesep\ignorespaces#1\@finalstrut\strutbox}% 512 | \color@endgroup}}% 513 | \fi 514 | \long\def\@mpfootnotetext#1{% 515 | \global\setbox\@mpfootins\vbox{% 516 | \unvbox\@mpfootins 517 | \reset@font\footnotesize 518 | \hsize\columnwidth 519 | \@parboxrestore 520 | \protected@edef\@currentlabel 521 | {\csname p@mpfootnote\endcsname\@thefnmark}% 522 | \color@begingroup\centering 523 | \@makefntext{% 524 | \rule\z@\footnotesep\ignorespaces#1\@finalstrut\strutbox}% 525 | \color@endgroup}} 526 | \def\@makefnmark{\hbox{\@textsuperscript{\normalfont\@thefnmark}}} 527 | \newif\if@ACM@newfonts 528 | \@ACM@newfontstrue 529 | \IfFileExists{libertine.sty}{}{\ClassWarning{\@classname}{You do not 530 | have libertine package installed. Please upgrade your 531 | TeX}\@ACM@newfontsfalse} 532 | \IfFileExists{zi4.sty}{}{\ClassWarning{\@classname}{You do not 533 | have zi4 package installed. Please upgrade your TeX}\@ACM@newfontsfalse} 534 | \IfFileExists{newtxmath.sty}{}{\ClassWarning{\@classname}{You do not 535 | have newtxmath package installed. Please upgrade your 536 | TeX}\@ACM@newfontsfalse} 537 | \if@ACM@newfonts 538 | \RequirePackage[tt=false]{libertine} 539 | \RequirePackage[varqu]{zi4} 540 | \RequirePackage[libertine]{newtxmath} 541 | \else 542 | \RequirePackage{textcomp} 543 | \fi 544 | \if@ACM@sigchiamode 545 | \renewcommand{\familydefault}{\sfdefault} 546 | \fi 547 | \RequirePackage{caption, float} 548 | \captionsetup[table]{position=top} 549 | \if@ACM@journal 550 | \captionsetup{labelfont={sf, small}, 551 | textfont={sf, small}, margin=\z@} 552 | \captionsetup[figure]{name={Fig.}} 553 | \else 554 | \captionsetup{labelfont={bf}, 555 | textfont={bf}, labelsep=colon, margin=\z@} 556 | \ifcase\ACM@format@nr 557 | \relax % manuscript 558 | \or % acmsmall 559 | \or % acmlarge 560 | \or % acmtog 561 | \or % sigconf 562 | \or % siggraph 563 | \captionsetup{textfont={it}} 564 | \or % sigplan 565 | \captionsetup{labelfont={bf}, 566 | textfont={normalfont}, labelsep=period, margin=\z@} 567 | \or % sigchi 568 | \captionsetup[figure]{labelfont={bf, small}, 569 | textfont={bf, small}} 570 | \or % sigchi-a 571 | \captionsetup[figure]{labelfont={bf, small}, 572 | textfont={bf, small}} 573 | \fi 574 | \fi 575 | \newfloat{sidebar}{}{sbar} 576 | \floatname{sidebar}{Sidebar} 577 | \renewenvironment{sidebar}{\Collect@Body\@sidebar}{} 578 | \long\def\@sidebar#1{\bgroup\captionsetup{type=sidebar}% 579 | \marginpar{\small#1}\egroup} 580 | \newenvironment{marginfigure}{\Collect@Body\@marginfigure}{} 581 | \long\def\@marginfigure#1{\bgroup\captionsetup{type=figure}% 582 | \marginpar{\centering\small#1}\egroup} 583 | \newenvironment{margintable}{\Collect@Body\@margintable}{} 584 | \long\def\@margintable#1{\bgroup\captionsetup{type=table}% 585 | \marginpar{\centering\small#1}\egroup} 586 | \newdimen\fulltextwidth 587 | \fulltextwidth=\dimexpr(\textwidth+\marginparwidth+\marginparsep) 588 | \if@ACM@sigchiamode 589 | \def\@dblfloat{\bgroup\columnwidth=\fulltextwidth 590 | \let\@endfloatbox\@endwidefloatbox 591 | \def\@fpsadddefault{\def\@fps{tp}}% 592 | \@float} 593 | \fi 594 | \if@ACM@sigchiamode 595 | \def\end@dblfloat{% 596 | \end@float\egroup} 597 | \fi 598 | \def\@endwidefloatbox{% 599 | \par\vskip\z@skip 600 | \@minipagefalse 601 | \outer@nobreak 602 | \egroup 603 | \color@endbox 604 | \global\setbox\@currbox=\vbox{\moveleft 605 | \dimexpr(\fulltextwidth-\textwidth)\box\@currbox}% 606 | \wd\@currbox=\textwidth 607 | } 608 | \ifcase\ACM@format@nr 609 | \relax % manuscript 610 | \or % acmsmall 611 | \or % acmlarge 612 | \or % acmtog 613 | \or % sigconf 614 | \or % siggraph 615 | \or % sigplan 616 | \def\labelenumi{\theenumi.} 617 | \def\labelenumii{\theenumii.} 618 | \def\labelenumiii{\theenumiii.} 619 | \def\labelenumiv{\theenumiv.} 620 | \or % sigchi 621 | \or % sigchi-a 622 | \fi 623 | \renewcommand{\descriptionlabel}[1]{\hspace\labelsep \upshape\bfseries #1} 624 | \renewenvironment{description}{\list{}{% 625 | \itemindent-12\p@ 626 | \labelwidth\z@ \let\makelabel\descriptionlabel}% 627 | }{ 628 | \endlist 629 | } 630 | \let\enddescription=\endlist % for efficiency 631 | \define@choicekey*+{ACM}{acmJournal}[\@journalCode\@journalCode@nr]{% 632 | CIE,% 633 | CSUR,% 634 | IMWUT,% 635 | JACM,% 636 | JDIQ,% 637 | JEA,% 638 | JERIC,% 639 | JETC,% 640 | JOCCH,% 641 | PACMPL,% 642 | TAAS,% 643 | TACCESS,% 644 | TACO,% 645 | TALG,% 646 | TALLIP,% 647 | TAP,% 648 | TCPS,% 649 | TEAC,% 650 | TECS,% 651 | TIIS,% 652 | TISSEC,% 653 | TIST,% 654 | TKDD,% 655 | TMIS,% 656 | TOCE,% 657 | TOCHI,% 658 | TOCL,% 659 | TOCS,% 660 | TOCT,% 661 | TODAES,% 662 | TODS,% 663 | TOG,% 664 | TOIS,% 665 | TOIT,% 666 | TOMACS,% 667 | TOMM,% 668 | TOMPECS,% 669 | TOMS,% 670 | TOPC,% 671 | TOPS,% 672 | TOPLAS,% 673 | TOS,% 674 | TOSEM,% 675 | TOSN,% 676 | TRETS,% 677 | TSAS,% 678 | TSC,% 679 | TSLP,% 680 | TWEB% 681 | }{% 682 | \ifcase\@journalCode@nr 683 | \relax % CIE 684 | \def\@journalName{ACM Computers in Entertainment}% 685 | \def\@journalNameShort{ACM Comput. Entertain.}% 686 | \def\@permissionCodeOne{1544-3574}% 687 | \or % CSUR 688 | \def\@journalName{ACM Computing Surveys}% 689 | \def\@journalNameShort{ACM Comput. Surv.}% 690 | \def\@permissionCodeOne{0360-0300}% 691 | \or % IMWUT 692 | \def\@journalName{PACM on Interactive, Mobile, Wearable and 693 | Ubiquitous Technologies}% 694 | \def\@journalNameShort{PACM Interact. Mob. Wearable Ubiquitous Technol.}% 695 | \def\@permissionCodeOne{2474-9567}% 696 | \or % JACM 697 | \def\@journalName{Journal of the ACM}% 698 | \def\@journalNameShort{J. ACM}% 699 | \def\@permissionCodeOne{0004-5411}% 700 | \or % JDIQ 701 | \def\@journalName{ACM Journal of Data and Information Quality}% 702 | \def\@journalNameShort{ACM J. Data Inform. Quality}% 703 | \def\@permissionCodeOne{1936-1955}% 704 | \or % JEA 705 | \def\@journalName{ACM Journal of Experimental Algorithmics}% 706 | \def\@journalNameShort{ACM J. Exp. Algor.}% 707 | \def\@permissionCodeOne{1084-6654}% 708 | \or % JERIC 709 | \def\@journalName{ACM Journal of Educational Resources in Computing}% 710 | \def\@journalNameShort{ACM J. Edu. Resources in Comput.}% 711 | \def\@permissionCodeOne{1073-0516}% 712 | \or % JETC 713 | \def\@journalName{ACM Journal on Emerging Technologies in Computing Systems}% 714 | \def\@journalNameShort{ACM J. Emerg. Technol. Comput. Syst.}% 715 | \def\@permissionCodeOne{1550-4832}% 716 | \or % JOCCH 717 | \def\@journalName{ACM Journal on Computing and Cultural Heritage}% 718 | \def\@journalName{ACM J. Comput. Cult. Herit.}% 719 | \or % PACMPL 720 | \def\@journalName{PACM on Programming Languages}% 721 | \def\@journalName{PACM Progr. Lang.}% 722 | \def\@permissionCodeOne{2475-1421}% 723 | \or % TAAS 724 | \def\@journalName{ACM Transactions on Autonomous and Adaptive Systems}% 725 | \def\@journalNameShort{ACM Trans. Autonom. Adapt. Syst.}% 726 | \def\@permissionCodeOne{1556-4665}% 727 | \or % TACCESS 728 | \def\@journalName{ACM Transactions on Accessible Computing}% 729 | \def\@journalNameShort{ACM Trans. Access. Comput.}% 730 | \def\@permissionCodeOne{1936-7228}% 731 | \or % TACO 732 | \def\@journalName{ACM Transactions on Architecture and Code Optimization}% 733 | \or % TALG 734 | \def\@journalName{ACM Transactions on Algorithms}% 735 | \def\@journalNameShort{ACM Trans. Algor.}% 736 | \def\@permissionCodeOne{1549-6325}% 737 | \or % TALLIP 738 | \def\@journalName{ACM Transactions on Asian and Low-Resource Language Information Processing}% 739 | \def\@journalNameShort{ACM Trans. Asian Low-Resour. Lang. Inf. Process.}% 740 | \def\@permissionCodeOne{2375-4699}% 741 | \or % TAP 742 | \def\@journalName{ACM Transactions on Applied Perception}% 743 | \or % TCPS 744 | \def\@journalName{ACM Transactions on Cyber-Physical Systems}% 745 | \or % TEAC 746 | \def\@journalName{ACM Transactions on Economics and Computation}% 747 | \or % TECS 748 | \def\@journalName{ACM Transactions on Embedded Computing Systems}% 749 | \def\@journalNameShort{ACM Trans. Embedd. Comput. Syst.}% 750 | \def\@permissionCodeOne{1539-9087}% 751 | \or % TIIS 752 | \def\@journalName{ACM Transactions on Interactive Intelligent Systems}% 753 | \def\@journalNameShort{ACM Trans. Interact. Intell. Syst.}% 754 | \def\@permissionCodeOne{2160-6455}% 755 | \or % TISSEC 756 | \def\@journalName{ACM Transactions on Information and System Security}% 757 | \def\@journalNameShort{ACM Trans. Info. Syst. Sec.}% 758 | \def\@permissionCodeOne{1094-9224}% 759 | \or % TIST 760 | \def\@journalName{ACM Transactions on Intelligent Systems and Technology}% 761 | \def\@journalNameShort{ACM Trans. Intell. Syst. Technol.}% 762 | \def\@permissionCodeOne{2157-6904}% 763 | \or % TKDD 764 | \def\@journalName{ACM Transactions on Knowledge Discovery from Data}% 765 | \def\@journalNameShort{ACM Trans. Knowl. Discov. Data.}% 766 | \def\@permissionCodeOne{1556-4681}% 767 | \or % TMIS 768 | \def\@journalName{ACM Transactions on Management Information Systems}% 769 | \def\@journalNameShort{ACM Trans. Manag. Inform. Syst.}% 770 | \def\@permissionCodeOne{2158-656X}% 771 | \or % TOCE 772 | \def\@journalName{ACM Transactions on Computing Education}% 773 | \def\@journalNameShort{ACM Trans. Comput. Educ.}% 774 | \def\@permissionCodeOne{1946-6226}% 775 | \or % TOCHI 776 | \def\@journalName{ACM Transactions on Computer-Human Interaction}% 777 | \def\@journalNameShort{ACM Trans. Comput.-Hum. Interact.}% 778 | \def\@permissionCodeOne{1073-0516}% 779 | \or % TOCL 780 | \def\@journalName{ACM Transactions on Computational Logic}% 781 | \def\@journalNameShort{ACM Trans. Comput. Logic}% 782 | \def\@permissionCodeOne{1529-3785}% 783 | \or % TOCS 784 | \def\@journalName{ACM Transactions on Computer Systems}% 785 | \def\@journalNameShort{ACM Trans. Comput. Syst.}% 786 | \def\@permissionCodeOne{0734-2071}% 787 | \or % TOCT 788 | \def\@journalName{ACM Transactions on Computation Theory}% 789 | \def\@journalNameShort{ACM Trans. Comput. Theory}% 790 | \def\@permissionCodeOne{1942-3454}% 791 | \or % TODAES 792 | \def\@journalName{ACM Transactions on Design Automation of Electronic Systems}% 793 | \def\@journalNameShort{ACM Trans. Des. Autom. Electron. Syst.}% 794 | \def\@permissionCodeOne{1084-4309}% 795 | \or % TODS 796 | \def\@journalName{ACM Transactions on Database Systems}% 797 | \def\@journalNameShort{ACM Trans. Datab. Syst.}% 798 | \def\@permissionCodeOne{0362-5915}% 799 | \or % TOG 800 | \def\@journalName{ACM Transactions on Graphics}% 801 | \def\@journalNameShort{ACM Trans. Graph.}% 802 | \def\@permissionCodeOne{0730-0301} 803 | \or % TOIS 804 | \def\@journalName{ACM Transactions on Information Systems}% 805 | \def\@journalName{ACM Transactions on Information Systems}% 806 | \def\@permissionCodeOne{1046-8188}% 807 | \or % TOIT 808 | \def\@journalName{ACM Transactions on Internet Technology}% 809 | \def\@journalNameShort{ACM Trans. Internet Technol.}% 810 | \def\@permissionCodeOne{1533-5399}% 811 | \or % TOMACS 812 | \def\@journalName{ACM Transactions on Modeling and Computer Simulation}% 813 | \def\@journalName{ACM Transactions on Modeling and Computer Simulation}% 814 | \def\@journalNameShort{ACM Trans. Model. Comput. Simul.}% 815 | \or % TOMM 816 | \def\@journalName{ACM Transactions on Multimedia Computing, Communications and Applications}% 817 | \def\@journalNameShort{ACM Trans. Multimedia Comput. Commun. Appl.}% 818 | \def\@permissionCodeOne{1551-6857}% 819 | \def\@permissionCodeTwo{0100}% 820 | \or % TOMPECS 821 | \def\@journalName{ACM Transactions on Modeling and Performance Evaluation of Computing Systems}% 822 | \def\@journalNameShort{ACM Trans. Model. Perform. Eval. Comput. Syst.}% 823 | \def\@permissionCodeOne{2376-3639}% 824 | \or % TOMS 825 | \def\@journalName{ACM Transactions on Mathematical Software}% 826 | \def\@journalNameShort{ACM Trans. Math. Softw.}% 827 | \def\@permissionCodeOne{0098-3500}% 828 | \or % TOPC 829 | \def\@journalName{ACM Transactions on Parallel Computing}% 830 | \def\@journalNameShort{ACM Trans. Parallel Comput.}% 831 | \def\@permissionCodeOne{1539-9087}% 832 | \or % TOPS 833 | \def\@journalName{ACM Transactions on Privacy and Security}% 834 | \def\@journalNameShort{ACM Trans. Priv. Sec.}% 835 | \def\@permissionCodeOne{2471-2566}% 836 | \or % TOPLAS 837 | \def\@journalName{ACM Transactions on Programming Languages and Systems}% 838 | \def\@journalNameShort{ACM Trans. Program. Lang. Syst.}% 839 | \def\@permissionCodeOne{0164-0925}% 840 | \or % TOS 841 | \def\@journalName{ACM Transactions on Storage}% 842 | \def\@journalNameShort{ACM Trans. Storage}% 843 | \def\@permissionCodeOne{1553-3077}% 844 | \or % TOSEM 845 | \def\@journalName{ACM Transactions on Software Engineering and Methodology}% 846 | \def\@journalNameShort{ACM Trans. Softw. Eng. Methodol.}% 847 | \def\@permissionCodeOne{1049-331X}% 848 | \or % TOSN 849 | \def\@journalName{ACM Transactions on Sensor Networks}% 850 | \def\@journalNameShort{ACM Trans. Sensor Netw.}% 851 | \def\@permissionCodeOne{1550-4859}% 852 | \or % TRETS 853 | \def\@journalName{ACM Transactions on Reconfigurable Technology and Systems}% 854 | \def\@journalNameShort{ACM Trans. Reconfig. Technol. Syst.}% 855 | \def\@permissionCodeOne{1936-7406}% 856 | \or % TSAS 857 | \def\@journalName{ACM Transactions on Spatial Algorithms and Systems}% 858 | \def\@journalNameShort{ACM Trans. Spatial Algorithms Syst.}% 859 | \def\@permissionCodeOne{2374-0353}% 860 | \or % TSC 861 | \def\@journalName{ACM Transactions on Social Computing}% 862 | \def\@journalNameShort{ACM Trans. Soc. Comput.}% 863 | \def\@permissionCodeOne{2469-7818}% 864 | \or % TSLP 865 | \def\@journalName{ACM Transactions on Speech and Language Processing}% 866 | \def\@journalNameShort{ACM Trans. Speech Lang. Process.}% 867 | \def\@permissionCodeOne{1550-4875}% 868 | \or % TWEB 869 | \def\@journalName{ACM Transactions on the Web}% 870 | \def\@journalNameShort{ACM Trans. Web}% 871 | \def\@permissionCodeOne{1559-1131}% 872 | \fi 873 | \ClassInfo{\@classname}{Using journal code \@journalCode}% 874 | }{% 875 | \ClassError{\@classname}{Incorrect journal #1}% 876 | }% 877 | \def\acmJournal#1{\setkeys{ACM}{acmJournal=#1}} 878 | \def\@journalCode@nr{0} 879 | \def\@journalName{}% 880 | \def\@journalNameShort{\@journalName}% 881 | \def\@permissionCodeOne{XXXX-XXXX}% 882 | \def\@permissionCodeTwo{}% 883 | \newcommand\acmConference[4][]{% 884 | \gdef\acmConference@shortname{#1}% 885 | \gdef\acmConference@name{#2}% 886 | \gdef\acmConference@date{#3}% 887 | \gdef\acmConference@venue{#4}% 888 | \ifx\acmConference@shortname\@empty 889 | \gdef\acmConference@shortname{#2}% 890 | \fi} 891 | \acmConference[Conference'17]{ACM Conference}{July 2017}{Washington, 892 | DC, USA} 893 | \def\subtitle#1{\def\@subtitle{#1}} 894 | \subtitle{} 895 | \newcount\num@authorgroups 896 | \num@authorgroups=0\relax 897 | \newif\if@insideauthorgroup 898 | \@insideauthorgroupfalse 899 | \renewcommand\author[2][]{% 900 | \if@insideauthorgroup\else 901 | \global\advance\num@authorgroups by 1\relax 902 | \global\@insideauthorgrouptrue 903 | \fi 904 | \ifx\addresses\@empty 905 | \if@ACM@anonymous 906 | \gdef\addresses{\@author{Anonymous Author(s)}}% 907 | \gdef\authors{Anonymous Author(s)}% 908 | \else 909 | \gdef\addresses{\@author{#2}}% 910 | \gdef\authors{#2}% 911 | \fi 912 | \else 913 | \if@ACM@anonymous\else 914 | \g@addto@macro\addresses{\and\@author{#2}}% 915 | \g@addto@macro\authors{\and#2}% 916 | \fi 917 | \fi 918 | \if@ACM@anonymous 919 | \ifx\shortauthors\@empty 920 | \gdef\shortauthors{Anon.}% 921 | \fi 922 | \else 923 | \def\@tempa{#1}% 924 | \ifx\@tempa\@empty 925 | \ifx\shortauthors\@empty 926 | \gdef\shortauthors{#2}% 927 | \else 928 | \g@addto@macro\shortauthors{\and#2}% 929 | \fi 930 | \else 931 | \ifx\shortauthors\@empty 932 | \gdef\shortauthors{#1}% 933 | \else 934 | \g@addto@macro\shortauthors{\and#1}% 935 | \fi 936 | \fi 937 | \fi} 938 | \newcommand{\affiliation}[2][]{% 939 | \global\@insideauthorgroupfalse 940 | \if@ACM@anonymous\else 941 | \g@addto@macro\addresses{\affiliation{#1}{#2}}% 942 | \fi} 943 | \renewcommand{\email}[2][]{% 944 | \if@ACM@anonymous\else 945 | \g@addto@macro\addresses{\email{#1}{#2}}% 946 | \fi} 947 | \let\orcid\@gobble 948 | \def\@titlenotes{} 949 | \def\titlenote#1{% 950 | \g@addto@macro\@title{\footnotemark}% 951 | \if@ACM@anonymous 952 | \g@addto@macro\@titlenotes{% 953 | \stepcounter{footnote}\footnotetext{Title note}}% 954 | \else 955 | \g@addto@macro\@titlenotes{\stepcounter{footnote}\footnotetext{#1}}% 956 | \fi} 957 | \def\@subtitlenotes{} 958 | \def\subtitlenote#1{% 959 | \g@addto@macro\@subtitle{\footnotemark}% 960 | \if@ACM@anonymous 961 | \g@addto@macro\@subtitlenotes{% 962 | \stepcounter{footnote}\footnotetext{Subtitle note}}% 963 | \else 964 | \g@addto@macro\@subtitlenotes{% 965 | \stepcounter{footnote}\footnotetext{#1}}% 966 | \fi} 967 | \def\@authornotes{} 968 | \def\authornote#1{% 969 | \if@ACM@anonymous\else 970 | \g@addto@macro\addresses{\@authornotemark} 971 | \g@addto@macro\@authornotes{% 972 | \stepcounter{footnote}\footnotetext{#1}}% 973 | \fi} 974 | \def\acmVolume#1{\def\@acmVolume{#1}} 975 | \acmVolume{1} 976 | \def\acmNumber#1{\def\@acmNumber{#1}} 977 | \acmNumber{1} 978 | \def\acmArticle#1{\def\@acmArticle{#1}} 979 | \acmArticle{1} 980 | \def\acmArticleSeq#1{\def\@acmArticleSeq{#1}} 981 | \acmArticleSeq{\@acmArticle} 982 | \def\acmYear#1{\def\@acmYear{#1}} 983 | \acmYear{2016} 984 | \def\acmMonth#1{\def\@acmMonth{#1}} 985 | \acmMonth{1} 986 | \def\@acmPubDate{\ifcase\@acmMonth\or 987 | January\or February\or March\or April\or May\or June\or 988 | July\or August\or September\or October\or November\or 989 | December\fi~\@acmYear} 990 | \def\acmPrice#1{\def\@acmPrice{#1}} 991 | \acmPrice{15.00} 992 | \def\acmISBN#1{\def\@acmISBN{#1}} 993 | \acmISBN{978-x-xxxx-xxxx-x/YY/MM} 994 | \def\acmDOI#1{\def\@acmDOI{#1}} 995 | \acmDOI{10.1145/nnnnnnn.nnnnnnn} 996 | \newif\if@ACM@badge 997 | \@ACM@badgefalse 998 | \newlength\@ACM@badge@width 999 | \setlength\@ACM@badge@width{5pc} 1000 | \newlength\@ACM@title@width 1001 | \newlength\@ACM@badge@skip 1002 | \setlength\@ACM@badge@skip{1pc} 1003 | \newcommand\acmBadgeR[2][]{\@ACM@badgetrue 1004 | \def\@acmBadgeR@url{#1}% 1005 | \def\@acmBadgeR@image{#2}} 1006 | \def\@acmBadgeR@url{} 1007 | \def\@acmBadgeR@image{} 1008 | \newcommand\acmBadgeL[2][]{\@ACM@badgetrue 1009 | \def\@acmBadgeL@url{#1}% 1010 | \def\@acmBadgeL@image{#2}} 1011 | \def\@acmBadgeL@url{} 1012 | \def\@acmBadgeL@image{} 1013 | \def\startPage#1{\def\@startPage{#1}} 1014 | \startPage{} 1015 | \def\terms#1{\def\@terms{#1}} 1016 | \terms{} 1017 | \def\keywords#1{\def\@keywords{#1}} 1018 | \keywords{} 1019 | \renewenvironment{abstract}{\Collect@Body\@saveabstract}{} 1020 | \long\def\@saveabstract#1{\long\gdef\@abstract{#1}} 1021 | \@saveabstract{} 1022 | \long\def\@lempty{} 1023 | \define@boolkey+{@ACM@topmatter@}[@ACM@]{printcss}[true]{% 1024 | \if@ACM@printcss 1025 | \ClassInfo{\@classname}{Printing CSS}% 1026 | \else 1027 | \ClassInfo{\@classname}{Suppressing CSS}% 1028 | \fi}{\ClassError{\@classname}{printcss must be true or false}} 1029 | \define@boolkey+{@ACM@topmatter@}[@ACM@]{printacmref}[true]{% 1030 | \if@ACM@printacmref 1031 | \ClassInfo{\@classname}{Printing bibformat}% 1032 | \else 1033 | \ClassInfo{\@classname}{Suppressing bibformat}% 1034 | \fi}{\ClassError{\@classname}{printacmref must be true or false}} 1035 | \define@boolkey+{@ACM@topmatter@}[@ACM@]{printfolios}[true]{% 1036 | \if@ACM@printfolios 1037 | \ClassInfo{\@classname}{Printing folios}% 1038 | \else 1039 | \ClassInfo{\@classname}{Suppressing folios}% 1040 | \fi}{\ClassError{\@classname}{printfolios must be true or false}} 1041 | \def\settopmatter#1{\setkeys{@ACM@topmatter@}{#1}} 1042 | \settopmatter{printcss=true, printacmref=true} 1043 | \if@ACM@manuscript 1044 | \settopmatter{printfolios=true} 1045 | \else 1046 | \if@ACM@journal 1047 | \settopmatter{printfolios=true} 1048 | \else 1049 | \settopmatter{printfolios=false} 1050 | \fi 1051 | \fi 1052 | \def\@received{} 1053 | \newcommand\received[2][]{\def\@tempa{#1}% 1054 | \ifx\@tempa\@empty 1055 | \ifx\@received\@empty 1056 | \gdef\@received{Received #2}% 1057 | \else 1058 | \g@addto@macro{\@received}{; revised #2}% 1059 | \fi 1060 | \else 1061 | \ifx\@received\@empty 1062 | \gdef\@received{#1 #2}% 1063 | \else 1064 | \g@addto@macro{\@received}{; #1 #2}% 1065 | \fi 1066 | \fi} 1067 | \AtEndDocument{% 1068 | \ifx\@received\@empty\else 1069 | \par\bigskip\noindent\small\normalfont\@received\par 1070 | \fi} 1071 | \RequirePackage{comment} 1072 | \excludecomment{CCSXML} 1073 | \let\@concepts\@empty 1074 | \newcommand\ccsdesc[2][100]{% 1075 | \ccsdesc@parse#1~#2~} 1076 | \def\ccsdesc@parse#1~#2~#3~{% 1077 | \expandafter\ifx\csname CCS@#2\endcsname\relax 1078 | \expandafter\gdef\csname CCS@#2\endcsname{\textbullet\textbf{#2} $\to$ }% 1079 | \g@addto@macro{\@concepts}{\csname CCS@#2\endcsname}\fi 1080 | \expandafter\g@addto@macro\expandafter{\csname CCS@#2\endcsname}{% 1081 | \ifnum#1>499\textbf{#3; }\else 1082 | \ifnum#1>299\textit{#3; }\else 1083 | #3; \fi\fi}} 1084 | \newif\if@printcopyright 1085 | \@printcopyrighttrue 1086 | \newif\if@printpermission 1087 | \@printpermissiontrue 1088 | \newif\if@acmowned 1089 | \@acmownedtrue 1090 | \define@choicekey*{ACM@}{acmcopyrightmode}[% 1091 | \acm@copyrightinput\acm@copyrightmode]{none,acmcopyright,acmlicensed,% 1092 | rightsretained,usgov,usgovmixed,cagov,cagovmixed,% 1093 | licensedusgovmixed,licensedcagovmixed,othergov,licensedothergov}{% 1094 | \@printpermissiontrue 1095 | \@printcopyrighttrue 1096 | \@acmownedtrue 1097 | \ifnum\acm@copyrightmode=0\relax % none 1098 | \@printpermissionfalse 1099 | \@printcopyrightfalse 1100 | \@acmownedfalse 1101 | \fi 1102 | \ifnum\acm@copyrightmode=2\relax % acmlicensed 1103 | \@acmownedfalse 1104 | \fi 1105 | \ifnum\acm@copyrightmode=3\relax % rightsretained 1106 | \@acmownedfalse 1107 | \fi 1108 | \ifnum\acm@copyrightmode=4\relax % usgov 1109 | \@printpermissiontrue 1110 | \@printcopyrightfalse 1111 | \@acmownedfalse 1112 | \fi 1113 | \ifnum\acm@copyrightmode=6\relax % cagov 1114 | \@acmownedfalse 1115 | \fi 1116 | \ifnum\acm@copyrightmode=8\relax % licensedusgovmixed 1117 | \@acmownedfalse 1118 | \fi 1119 | \ifnum\acm@copyrightmode=9\relax % licensedcagovmixed 1120 | \@acmownedfalse 1121 | \fi 1122 | \ifnum\acm@copyrightmode=10\relax % othergov 1123 | \@acmownedtrue 1124 | \fi 1125 | \ifnum\acm@copyrightmode=11\relax % licensedothergov 1126 | \@acmownedfalse 1127 | \fi} 1128 | \def\setcopyright#1{\setkeys{ACM@}{acmcopyrightmode=#1}} 1129 | \setcopyright{acmcopyright} 1130 | \def\@copyrightowner{% 1131 | \ifcase\acm@copyrightmode\relax % none 1132 | \or % acmcopyright 1133 | ACM\@. 1134 | \or % acmlicensed 1135 | Copyright held by the owner/author(s). Publication rights licensed to 1136 | ACM\@. 1137 | \or % rightsretained 1138 | Copyright held by the owner/author(s). 1139 | \or % usgov 1140 | \or % usgovmixed 1141 | ACM\@. 1142 | \or % cagov 1143 | Crown in Right of Canada. 1144 | \or %cagovmixed 1145 | ACM\@. 1146 | \or %licensedusgovmixed 1147 | Copyright held by the owner/author(s). Publication rights licensed to 1148 | ACM\@. 1149 | \or %licensedcagovmixed 1150 | Copyright held by the owner/author(s). Publication rights licensed to 1151 | ACM\@. 1152 | \or % othergov 1153 | ACM\@. 1154 | \or % licensedothergov 1155 | Copyright held by the owner/author(s). Publication rights licensed to 1156 | ACM\@. 1157 | \fi} 1158 | \def\@formatdoi#1{\url{http://dx.doi.org/#1}} 1159 | \def\@copyrightpermission{% 1160 | \ifcase\acm@copyrightmode\relax % none 1161 | \or % acmcopyright 1162 | Permission to make digital or hard copies of all or part of this 1163 | work for personal or classroom use is granted without fee provided 1164 | that copies are not made or distributed for profit or commercial 1165 | advantage and that copies bear this notice and the full citation on 1166 | the first page. Copyrights for components of this work owned by 1167 | others than ACM must be honored. Abstracting with credit is 1168 | permitted. To copy otherwise, or republish, to post on servers or to 1169 | redistribute to lists, requires prior specific permission 1170 | and\hspace*{.5pt}/or a fee. Request permissions from 1171 | permissions@acm.org. 1172 | \or % acmlicensed 1173 | Permission to make digital or hard copies of all or part of this 1174 | work for personal or classroom use is granted without fee provided 1175 | that copies are not made or distributed for profit or commercial 1176 | advantage and that copies bear this notice and the full citation on 1177 | the first page. Copyrights for components of this work owned by 1178 | others than the author(s) must be honored. Abstracting with credit 1179 | is permitted. To copy otherwise, or republish, to post on servers 1180 | or to redistribute to lists, requires prior specific permission 1181 | and\hspace*{.5pt}/or a fee. Request permissions from 1182 | permissions@acm.org. 1183 | \or % rightsretained 1184 | Permission to make digital or hard copies of part or all of this work 1185 | for personal or classroom use is granted without fee provided that 1186 | copies are not made or distributed for profit or commercial advantage 1187 | and that copies bear this notice and the full citation on the first 1188 | page. Copyrights for third-party components of this work must be 1189 | honored. For all other uses, contact the 1190 | owner\hspace*{.5pt}/author(s). 1191 | \or % usgov 1192 | This paper is authored by an employee(s) of the United States 1193 | Government and is in the public domain. Non-exclusive copying or 1194 | redistribution is allowed, provided that the article citation is 1195 | given and the authors and agency are clearly identified as its 1196 | source. 1197 | \or % usgovmixed 1198 | ACM acknowledges that this contribution was authored or co-authored 1199 | by an employee, or contractor of the national government. As such, 1200 | the Government retains a nonexclusive, royalty-free right to 1201 | publish or reproduce this article, or to allow others to do so, for 1202 | Government purposes only. Permission to make digital or hard copies 1203 | for personal or classroom use is granted. Copies must bear this 1204 | notice and the full citation on the first page. Copyrights for 1205 | components of this work owned by others than ACM must be 1206 | honored. To copy otherwise, distribute, republish, or post, 1207 | requires prior specific permission and\hspace*{.5pt}/or a 1208 | fee. Request permissions from permissions@acm.org. 1209 | \or % cagov 1210 | This article was authored by employees of the Government of Canada. 1211 | As such, the Canadian government retains all interest in the 1212 | copyright to this work and grants to ACM a nonexclusive, 1213 | royalty-free right to publish or reproduce this article, or to allow 1214 | others to do so, provided that clear attribution is given both to 1215 | the authors and the Canadian government agency employing them. 1216 | Permission to make digital or hard copies for personal or classroom 1217 | use is granted. Copies must bear this notice and the full citation 1218 | on the first page. Copyrights for components of this work owned by 1219 | others than the Canadain Government must be honored. To copy 1220 | otherwise, distribute, republish, or post, requires prior specific 1221 | permission and\hspace*{.5pt}/or a fee. Request permissions from 1222 | permissions@acm.org. 1223 | \or % cagovmixed 1224 | ACM acknowledges that this contribution was co-authored by an 1225 | affiliate of the national government of Canada. As such, the Crown 1226 | in Right of Canada retains an equal interest in the copyright. 1227 | Reprints must include clear attribution to ACM and the author's 1228 | government agency affiliation. Permission to make digital or hard 1229 | copies for personal or classroom use is granted. Copies must bear 1230 | this notice and the full citation on the first page. Copyrights for 1231 | components of this work owned by others than ACM must be honored. 1232 | To copy otherwise, distribute, republish, or post, requires prior 1233 | specific permission and\hspace*{.5pt}/or a fee. Request permissions 1234 | from permissions@acm.org. 1235 | \or % licensedusgovmixed 1236 | Publication rights licensed to ACM\@. ACM acknowledges that this 1237 | contribution was authored or co-authored by an employee, contractor 1238 | or affiliate of the United States government. As such, the 1239 | Government retains a nonexclusive, royalty-free right to publish or 1240 | reproduce this article, or to allow others to do so, for Government 1241 | purposes only. 1242 | \or % licensedcagovmixed 1243 | Publication rights licensed to ACM\@. ACM acknowledges that this 1244 | contribution was authored or co-authored by an employee, contractor 1245 | or affiliate of the national government of Canada. As such, the 1246 | Government retains a nonexclusive, royalty-free right to publish or 1247 | reproduce this article, or to allow others to do so, for Government 1248 | purposes only. 1249 | \or % othergov 1250 | ACM acknowledges that this contribution was authored or co-authored 1251 | by an employee, contractor or affiliate of a national government. As 1252 | such, the Government retains a nonexclusive, royalty-free right to 1253 | publish or reproduce this article, or to allow others to do so, for 1254 | Government purposes only. 1255 | \or % licensedothergov 1256 | Publication rights licensed to ACM\@. ACM acknowledges that this 1257 | contribution was authored or co-authored by an employee, contractor 1258 | or affiliate of a national government. As such, the Government 1259 | retains a nonexclusive, royalty-free right to publish or reproduce 1260 | this article, or to allow others to do so, for Government purposes 1261 | only. 1262 | \fi} 1263 | \def\copyrightyear#1{\def\@copyrightyear{#1}} 1264 | \copyrightyear{\@acmYear} 1265 | \def\@teaserfigures{} 1266 | \newenvironment{teaserfigure}{\Collect@Body\@saveteaser}{} 1267 | \long\def\@saveteaser#1{\g@addto@macro\@teaserfigures{\@teaser{#1}}} 1268 | \renewcommand{\thanks}[1]{% 1269 | \@ifnotempty{#1}{% 1270 | \if@ACM@anonymous 1271 | \g@addto@macro\thankses{\thanks{A note}}% 1272 | \else 1273 | \g@addto@macro\thankses{\thanks{#1}}% 1274 | \fi}} 1275 | \newbox\mktitle@bx 1276 | \def\maketitle{% 1277 | \if@ACM@anonymous 1278 | % Anonymize omission of \author-s 1279 | \ifnum\num@authorgroups=0\author{}\fi 1280 | \fi 1281 | \begingroup 1282 | \let\@footnotemark\@footnotemark@nolink 1283 | \let\@footnotetext\@footnotetext@nolink 1284 | \renewcommand\thefootnote{\@fnsymbol\c@footnote}% 1285 | \@topnum\z@ % this prevents figures from falling at the top of page 1286 | % 1 1287 | \hsize=\textwidth 1288 | \def\@makefnmark{\hbox{\@textsuperscript{\@thefnmark}}}% 1289 | \@mktitle\if@ACM@sigchiamode\else\@mkauthors\fi\@mkteasers 1290 | \@printtopmatter 1291 | \if@ACM@sigchiamode\@mkauthors\fi 1292 | \setcounter{footnote}{0}% 1293 | \def\@makefnmark{\hbox{\@textsuperscript{\normalfont\@thefnmark}}}% 1294 | \@titlenotes 1295 | \@subtitlenotes 1296 | \@authornotes 1297 | \let\@makefnmark\relax \let\@thefnmark\relax 1298 | \let\@makefntext\noindent 1299 | \ifx\@empty\thankses\else 1300 | \footnotetextcopyrightpermission{% 1301 | \def\par{\let\par\@par}\parindent\z@\@setthanks}% 1302 | \fi 1303 | \footnotetextcopyrightpermission{\parindent\z@\parskip0.1\baselineskip 1304 | \if@ACM@authorversion\else 1305 | \if@printpermission\@copyrightpermission\par\fi 1306 | \fi 1307 | \if@ACM@manuscript\else 1308 | \if@ACM@journal\else % Print the conference short name 1309 | {\itshape \acmConference@shortname, \acmConference@venue}\par 1310 | \fi 1311 | \fi 1312 | \if@printcopyright 1313 | \copyright\ \@copyrightyear\ \@copyrightowner\ 1314 | \else 1315 | \@copyrightyear.\ 1316 | \fi 1317 | \if@ACM@manuscript 1318 | Manuscript submitted to ACM\\ 1319 | \else 1320 | \if@ACM@authorversion 1321 | This is the author's version of the work. It is posted here for 1322 | your personal use. Not for redistribution. The definitive Version 1323 | of Record was published in 1324 | \if@ACM@journal 1325 | \emph{\@journalName}% 1326 | \else 1327 | \emph{Proceedings of \acmConference@name, \acmConference@date}% 1328 | \fi 1329 | \ifx\@acmDOI\@empty 1330 | . 1331 | \else 1332 | , \@formatdoi{\@acmDOI}. 1333 | \fi\\ 1334 | \else 1335 | \if@ACM@journal 1336 | \@permissionCodeOne/\@acmYear/\@acmMonth-ART\@acmArticle\ 1337 | \$\@acmPrice\\ 1338 | DOI: \nolinkurl{\@acmDOI}% 1339 | \else % Conference 1340 | \@acmISBN 1341 | \ifx\@acmPrice\@empty.\else\dots\$\@acmPrice\fi\\ 1342 | DOI: \nolinkurl{\@acmDOI}% 1343 | \fi 1344 | \fi 1345 | \fi}% 1346 | \endgroup 1347 | \setcounter{footnote}{0}% 1348 | \@mkabstract 1349 | \if@ACM@printcss 1350 | \ifx\@concepts\@empty\else\bgroup 1351 | {\@specialsection{CCS Concepts}% 1352 | \@concepts\par}\egroup 1353 | \fi 1354 | \fi 1355 | \if\@terms\@empty\else\bgroup 1356 | {\@specialsection{General Terms}% 1357 | \@terms\par}\egroup 1358 | \fi 1359 | \ifx\@keywords\@empty\else\bgroup 1360 | {\if@ACM@journal 1361 | \@specialsection{Additional Key Words and Phrases}% 1362 | \else 1363 | \@specialsection{Keywords}% 1364 | \fi 1365 | \@keywords}\par\egroup 1366 | \fi 1367 | \andify\authors 1368 | \andify\shortauthors 1369 | \global\let\authors=\authors 1370 | \global\let\shortauthors=\shortauthors 1371 | \if@ACM@printacmref 1372 | \@mkbibcitation 1373 | \fi 1374 | \hypersetup{pdfauthor={\authors}, 1375 | pdftitle={\@title}, pdfkeywords={\@concepts}}% 1376 | \@printendtopmatter 1377 | \@afterindentfalse 1378 | \@afterheading 1379 | } 1380 | \def\@specialsection#1{% 1381 | \ifcase\ACM@format@nr 1382 | \relax % manuscript 1383 | \par\medskip\small\noindent#1: % 1384 | \or % acmsmall 1385 | \par\medskip\small\noindent#1: % 1386 | \or % acmlarge 1387 | \par\medskip\small\noindent#1: % 1388 | \or % acmtog 1389 | \par\medskip\small\noindent#1: % 1390 | \or % sigconf 1391 | \section*{#1}% 1392 | \or % siggraph 1393 | \section*{#1}% 1394 | \or % sigplan 1395 | \paragraph*{#1}% 1396 | \or % sigchi 1397 | \section*{#1}% 1398 | \or % sigchi-a 1399 | \section*{#1}% 1400 | \fi} 1401 | \def\@printtopmatter{% 1402 | \ifx\@startPage\@empty 1403 | \gdef\@startPage{1}% 1404 | \else 1405 | \setcounter{page}{\@startPage}% 1406 | \fi 1407 | \thispagestyle{firstpagestyle}% 1408 | \noindent 1409 | \ifcase\ACM@format@nr 1410 | \relax % manuscript 1411 | \box\mktitle@bx\par 1412 | \noindent\hrulefill\par 1413 | \or % acmsmall 1414 | \box\mktitle@bx\par 1415 | \noindent\hrulefill\par 1416 | \or % acmlarge 1417 | \box\mktitle@bx\par 1418 | \noindent\hrulefill\par 1419 | \or % acmtog 1420 | \twocolumn[\box\mktitle@bx]% 1421 | \or % sigconf 1422 | \twocolumn[\box\mktitle@bx]% 1423 | \or % siggraph 1424 | \twocolumn[\box\mktitle@bx]% 1425 | \or % sigplan 1426 | \twocolumn[\box\mktitle@bx]% 1427 | \or % sigchi 1428 | \twocolumn[\box\mktitle@bx]% 1429 | \or % sigchi-a 1430 | \par\box\mktitle@bx\par\bigskip 1431 | \if@ACM@badge 1432 | \marginpar{\noindent 1433 | \ifx\@acmBadgeL@image\@empty\else 1434 | \href{\@acmBadgeL@url}{% 1435 | \includegraphics[width=\@ACM@badge@width]{\@acmBadgeL@image}}% 1436 | \hskip\@ACM@badge@skip 1437 | \fi 1438 | \ifx\@acmBadgeR@image\@empty\else 1439 | \href{\@acmBadgeR@url}{% 1440 | \includegraphics[width=\@ACM@badge@width]{\@acmBadgeR@image}}% 1441 | \fi}% 1442 | \fi 1443 | \fi 1444 | } 1445 | \def\@mktitle{% 1446 | \ifcase\ACM@format@nr 1447 | \relax % manuscript 1448 | \@mktitle@i 1449 | \or % acmsmall 1450 | \@mktitle@i 1451 | \or % acmlarge 1452 | \@mktitle@i 1453 | \or % acmtog 1454 | \@mktitle@i 1455 | \or % sigconf 1456 | \@mktitle@iii 1457 | \or % siggraph 1458 | \@mktitle@iii 1459 | \or % sigplan 1460 | \@mktitle@iii 1461 | \or % sigchi 1462 | \@mktitle@iii 1463 | \or % sigchi-a 1464 | \@mktitle@iv 1465 | \fi 1466 | } 1467 | \def\@titlefont{% 1468 | \ifcase\ACM@format@nr 1469 | \relax % manuscript 1470 | \LARGE\bfseries\sffamily 1471 | \or % acmsmall 1472 | \LARGE\bfseries\sffamily 1473 | \or % acmlarge 1474 | \LARGE\bfseries\sffamily 1475 | \or % acmtog 1476 | \Huge\sffamily 1477 | \or % sigconf 1478 | \Huge\sffamily\bfseries 1479 | \or % siggraph 1480 | \Huge\sffamily\bfseries 1481 | \or % sigplan 1482 | \Huge\bfseries 1483 | \or % sigchi 1484 | \Huge\sffamily\bfseries 1485 | \or % sigchi-a 1486 | \Huge\bfseries 1487 | \fi} 1488 | \def\@subtitlefont{% 1489 | \ifcase\ACM@format@nr 1490 | \relax % manuscript 1491 | \mdseries 1492 | \or % acmsmall 1493 | \mdseries 1494 | \or % acmlarge 1495 | \mdseries 1496 | \or % acmtog 1497 | \LARGE 1498 | \or % sigconf 1499 | \LARGE\mdseries 1500 | \or % siggraph 1501 | \LARGE\mdseries 1502 | \or % sigplan 1503 | \LARGE\mdseries 1504 | \or % sigchi 1505 | \LARGE\mdseries 1506 | \or % sigchi-a 1507 | \mdseries 1508 | \fi} 1509 | \def\@mktitle@i{\hsize=\textwidth 1510 | \@ACM@title@width=\hsize 1511 | \ifx\@acmBadgeL@image\@empty\else 1512 | \advance\@ACM@title@width by -\@ACM@badge@width 1513 | \advance\@ACM@title@width by -\@ACM@badge@skip 1514 | \fi 1515 | \ifx\@acmBadgeR@image\@empty\else 1516 | \advance\@ACM@title@width by -\@ACM@badge@width 1517 | \advance\@ACM@title@width by -\@ACM@badge@skip 1518 | \fi 1519 | \setbox\mktitle@bx=\vbox{\noindent\@titlefont 1520 | \ifx\@acmBadgeL@image\@empty\else 1521 | \raisebox{-.5\baselineskip}[\z@][\z@]{\href{\@acmBadgeL@url}{% 1522 | \includegraphics[width=\@ACM@badge@width]{\@acmBadgeL@image}}}% 1523 | \hskip\@ACM@badge@skip 1524 | \fi 1525 | \parbox[t]{\@ACM@title@width}{\raggedright 1526 | \@titlefont\noindent 1527 | \@title 1528 | \ifx\@subtitle\@empty\else 1529 | \par\noindent{\@subtitlefont\@subtitle} 1530 | \fi}% 1531 | \ifx\@acmBadgeR@image\@empty\else 1532 | \hskip\@ACM@badge@skip 1533 | \raisebox{-.5\baselineskip}[\z@][\z@]{\href{\@acmBadgeR@url}{% 1534 | \includegraphics[width=\@ACM@badge@width]{\@acmBadgeR@image}}}% 1535 | \fi 1536 | \par\bigskip}}% 1537 | \def\@mktitle@iii{\hsize=\textwidth 1538 | \setbox\mktitle@bx=\vbox{\@titlefont\centering 1539 | \@ACM@title@width=\hsize 1540 | \if@ACM@badge 1541 | \advance\@ACM@title@width by -2\@ACM@badge@width 1542 | \advance\@ACM@title@width by -2\@ACM@badge@skip 1543 | \parbox[b]{\@ACM@badge@width}{\strut 1544 | \ifx\@acmBadgeL@image\@empty\else 1545 | \raisebox{-.5\baselineskip}[\z@][\z@]{\href{\@acmBadgeL@url}{% 1546 | \includegraphics[width=\@ACM@badge@width]{\@acmBadgeL@image}}}% 1547 | \fi}% 1548 | \hskip\@ACM@badge@skip 1549 | \fi 1550 | \parbox[t]{\@ACM@title@width}{\centering\@titlefont 1551 | \@title 1552 | \ifx\@subtitle\@empty\else 1553 | \par\noindent{\@subtitlefont\@subtitle} 1554 | \fi 1555 | }% 1556 | \if@ACM@badge 1557 | \hskip\@ACM@badge@skip 1558 | \parbox[b]{\@ACM@badge@width}{\strut 1559 | \ifx\@acmBadgeR@image\@empty\else 1560 | \raisebox{-.5\baselineskip}[\z@][\z@]{\href{\@acmBadgeR@url}{% 1561 | \includegraphics[width=\@ACM@badge@width]{\@acmBadgeR@image}}}% 1562 | \fi}% 1563 | \fi 1564 | \par\bigskip}}% 1565 | \def\@mktitle@iv{\hsize=\textwidth 1566 | \setbox\mktitle@bx=\vbox{\raggedright\leftskip5pc\@titlefont 1567 | \noindent\leavevmode\leaders\hrule height 2pt\hfill\kern0pt\par 1568 | \noindent\@title 1569 | \ifx\@subtitle\@empty\else 1570 | \par\noindent\@subtitlefont\@subtitle 1571 | \fi 1572 | \par\bigskip}}% 1573 | \newbox\@ACM@commabox 1574 | \def\@ACM@addtoaddress#1{% 1575 | \ifvmode\else 1576 | \setbox\@ACM@commabox=\hbox{, }% 1577 | \unskip\cleaders\copy\@ACM@commabox\hskip\wd\@ACM@commabox 1578 | \fi 1579 | #1} 1580 | \if@ACM@journal 1581 | \let\position\@gobble 1582 | \def\institution#1{#1\ignorespaces}% 1583 | \let\department\@gobble 1584 | \let\streetaddress\@gobble 1585 | \let\city\@gobble 1586 | \let\state\@gobble 1587 | \let\postcode\@gobble 1588 | \let\country\@gobble 1589 | \else 1590 | \def\position#1{#1\par}% 1591 | \def\institution#1{#1\par}% 1592 | \def\department#1{#1\par}% 1593 | \def\streetaddress#1{#1\par}% 1594 | \let\city\@ACM@addtoaddress 1595 | \let\state\@ACM@addtoaddress 1596 | \def\postcode#1{\unskip\space#1}% 1597 | \let\country\@ACM@addtoaddress 1598 | \fi 1599 | \def\@mkauthors{\begingroup 1600 | \hsize=\textwidth 1601 | \ifcase\ACM@format@nr 1602 | \relax % manuscript 1603 | \@mkauthors@i 1604 | \or % acmsmall 1605 | \@mkauthors@i 1606 | \or % acmlarge 1607 | \@mkauthors@i 1608 | \or % acmtog 1609 | \@mkauthors@i 1610 | \or % sigconf 1611 | \@mkauthors@iii 1612 | \or % siggraph 1613 | \@mkauthors@iii 1614 | \or % sigplan 1615 | \@mkauthors@iii 1616 | \or % sigchi 1617 | \@mkauthors@iii 1618 | \or % sigchi-a 1619 | \@mkauthors@iv 1620 | \fi 1621 | \endgroup 1622 | } 1623 | \def\@authorfont{\Large\sffamily} 1624 | \def\@affiliationfont{\normalsize\normalfont} 1625 | \ifcase\ACM@format@nr 1626 | \relax % manuscript 1627 | \or % acmsmall 1628 | \def\@authorfont{\large\sffamily} 1629 | \def\@affiliationfont{\small\normalfont} 1630 | \or % acmlarge 1631 | \or % acmtog 1632 | \def\@authorfont{\LARGE\sffamily} 1633 | \def\@affiliationfont{\large} 1634 | \or % sigconf 1635 | \def\@authorfont{\LARGE} 1636 | \def\@affiliationfont{\large} 1637 | \or % siggraph 1638 | \def\@authorfont{\normalsize\normalfont} 1639 | \def\@affiliationfont{\normalsize\normalfont} 1640 | \or % sigplan 1641 | \def\@authorfont{\Large\normalfont} 1642 | \def\@affiliationfont{\normalsize\normalfont} 1643 | \or % sigchi 1644 | \def\@authorfont{\bfseries} 1645 | \def\@affiliationfont{\mdseries} 1646 | \or % sigchi-a 1647 | \def\@authorfont{\bfseries} 1648 | \def\@affiliationfont{\mdseries} 1649 | \fi 1650 | \def\@typeset@author@line{% 1651 | \andify\@currentauthors\par\noindent 1652 | \@currentauthors\def\@currentauthors{}% 1653 | \ifx\@currentaffiliations\@empty\else 1654 | \andify\@currentaffiliations 1655 | \unskip, {\@currentaffiliations}\par 1656 | \fi 1657 | \def\@currentaffiliations{}} 1658 | \def\@mkauthors@i{% 1659 | \def\@currentauthors{}% 1660 | \def\@currentaffiliations{}% 1661 | \global\let\and\@typeset@author@line 1662 | \def\@author##1{% 1663 | \ifx\@currentauthors\@empty 1664 | \gdef\@currentauthors{\@authorfont\MakeUppercase{##1}}% 1665 | \else 1666 | \g@addto@macro{\@currentauthors}{\and\MakeUppercase{##1}}% 1667 | \fi 1668 | \gdef\and{}}% 1669 | \def\email##1##2{}% 1670 | \def\affiliation##1##2{% 1671 | \def\@tempa{##2}\ifx\@tempa\@empty\else 1672 | \ifx\@currentaffiliations\@empty 1673 | \gdef\@currentaffiliations{\@affiliationfont##2}% 1674 | \else 1675 | \g@addto@macro{\@currentaffiliations}{\and##2}% 1676 | \fi 1677 | \fi 1678 | \global\let\and\@typeset@author@line} 1679 | \global\setbox\mktitle@bx=\vbox{\noindent\box\mktitle@bx\par\medskip 1680 | \noindent\addresses\@typeset@author@line 1681 | \par\medskip}% 1682 | } 1683 | \newbox\author@bx 1684 | \newdimen\author@bx@wd 1685 | \newskip\author@bx@sep 1686 | \author@bx@sep=1pc\relax 1687 | \def\@typeset@author@bx{\bgroup\hsize=\author@bx@wd\def\and{\par}% 1688 | \global\setbox\author@bx=\vtop{\if@ACM@sigchiamode\else\centering\fi 1689 | \@authorfont\@currentauthors\par\@affiliationfont 1690 | \@currentaffiliation}\egroup 1691 | \box\author@bx\hspace{\author@bx@sep}% 1692 | \gdef\@currentauthors{}% 1693 | \gdef\@currentaffiliation{}} 1694 | \def\@mkauthors@iii{% 1695 | \author@bx@wd=\textwidth\relax 1696 | \advance\author@bx@wd by -\author@bx@sep\relax 1697 | \ifcase\num@authorgroups 1698 | \relax % 0? 1699 | \or % 1=one author per row 1700 | \or % 2=two authors per row 1701 | \divide\author@bx@wd by \num@authorgroups\relax 1702 | \or % 3=three authors per row 1703 | \divide\author@bx@wd by \num@authorgroups\relax 1704 | \or % 4=two authors per row (!) 1705 | \divide\author@bx@wd by 2\relax 1706 | \else % three authors per row 1707 | \divide\author@bx@wd by 3\relax 1708 | \fi 1709 | \advance\author@bx@wd by -\author@bx@sep\relax 1710 | \gdef\@currentauthors{}% 1711 | \gdef\@currentaffiliation{}% 1712 | \def\@author##1{\ifx\@currentauthors\@empty 1713 | \gdef\@currentauthors{\par##1}% 1714 | \else 1715 | \g@addto@macro\@currentauthors{\par##1}% 1716 | \fi 1717 | \gdef\and{}}% 1718 | \def\email##1##2{\ifx\@currentaffiliation\@empty 1719 | \gdef\@currentaffiliation{\nolinkurl{##2}}% 1720 | \else 1721 | \g@addto@macro\@currentaffiliation{\par\nolinkurl{##2}}% 1722 | \fi}% 1723 | \def\affiliation##1##2{\ifx\@currentaffiliation\@empty 1724 | \gdef\@currentaffiliation{##2}% 1725 | \else 1726 | \g@addto@macro\@currentaffiliation{\par##2}% 1727 | \fi 1728 | \global\let\and\@typeset@author@bx 1729 | }% 1730 | \hsize=\textwidth 1731 | \global\setbox\mktitle@bx=\vbox{\noindent 1732 | \box\mktitle@bx\par\medskip\leavevmode 1733 | \lineskip=1pc\relax\centering\hspace*{-1em}% 1734 | \addresses\let\and\@typeset@author@bx\and\par\bigskip}} 1735 | \def\@mkauthors@iv{% 1736 | \author@bx@wd=\columnwidth\relax 1737 | \advance\author@bx@wd by -\author@bx@sep\relax 1738 | \ifcase\num@authorgroups 1739 | \relax % 0? 1740 | \or % 1=one author per row 1741 | \else % 2=two authors per row 1742 | \divide\author@bx@wd by 2\relax 1743 | \fi 1744 | \advance\author@bx@wd by -\author@bx@sep\relax 1745 | \gdef\@currentauthors{}% 1746 | \gdef\@currentaffiliation{}% 1747 | \def\@author##1{\ifx\@currentauthors\@empty 1748 | \gdef\@currentauthors{\par##1}% 1749 | \else 1750 | \g@addto@macro\@currentauthors{\par##1}% 1751 | \fi 1752 | \gdef\and{}}% 1753 | \def\email##1##2{\ifx\@currentaffiliation\@empty 1754 | \gdef\@currentaffiliation{\nolinkurl{##2}}% 1755 | \else 1756 | \g@addto@macro\@currentaffiliation{\par\nolinkurl{##2}}% 1757 | \fi}% 1758 | \def\affiliation##1##2{\ifx\@currentaffiliation\@empty 1759 | \gdef\@currentaffiliation{##2}% 1760 | \else 1761 | \g@addto@macro\@currentaffiliation{\par##2}% 1762 | \fi 1763 | \global\let\and\@typeset@author@bx}% 1764 | \bgroup\hsize=\columnwidth 1765 | \par\raggedright\leftskip=\z@ 1766 | \lineskip=1pc\noindent 1767 | \addresses\let\and\@typeset@author@bx\and\par\bigskip\egroup} 1768 | \def\@authornotemark{\g@addto@macro\@currentauthors{\footnotemark}} 1769 | \def\@mkteasers{% 1770 | \ifx\@teaserfigures\@empty\else 1771 | \def\@teaser##1{\par\bigskip\bgroup 1772 | \captionsetup{type=figure}##1\egroup\par} 1773 | \global\setbox\mktitle@bx=\vbox{\noindent\box\mktitle@bx\par 1774 | \noindent\@teaserfigures\par\medskip}% 1775 | \fi} 1776 | \def\@setaddresses{} 1777 | \def\@mkabstract{\bgroup 1778 | \ifx\@abstract\@lempty\else 1779 | {\if@ACM@journal 1780 | \small\noindent 1781 | \else 1782 | \section*{Abstract}% 1783 | \fi 1784 | \phantomsection\addcontentsline{toc}{section}{Abstract}% 1785 | \ignorespaces\@abstract\par}% 1786 | \fi\egroup} 1787 | \def\@mkbibcitation{\bgroup 1788 | \def\footnotemark{}% 1789 | \par\medskip\small\noindent{\bfseries ACM Reference format:}\par\nobreak 1790 | \noindent\authors. \@acmYear. \@title. 1791 | \if@ACM@journal 1792 | \textit{\@journalNameShort} 1793 | \@acmVolume, \@acmNumber, Article~\@acmArticle\ (\@acmPubDate), 1794 | \ref{TotPages}~pages. 1795 | \else 1796 | In \textit{Proceedings of \acmConference@name, \acmConference@venue, 1797 | \acmConference@date 1798 | \ifx\acmConference@name\acmConference@shortname\else 1799 | \ (\acmConference@shortname)\fi 1800 | ,} \ref{TotPages}~pages. 1801 | \fi\par 1802 | \noindent DOI: \nolinkurl{\@acmDOI} 1803 | \par\egroup} 1804 | \def\@printendtopmatter{\par\medskip 1805 | \ifcase\ACM@format@nr 1806 | \relax % manuscript 1807 | \noindent\hrulefill\par\medskip 1808 | \or % acmsmall 1809 | \noindent\hrulefill\par\medskip 1810 | \or % acmlarge 1811 | \noindent\hrulefill\par\medskip 1812 | \or % acmtog 1813 | \par\bigskip 1814 | \or % sigconf 1815 | \par\bigskip 1816 | \or % siggraph 1817 | \par\bigskip 1818 | \or % sigplan 1819 | \par\bigskip 1820 | \or % sigchi 1821 | \par\bigskip 1822 | \or % sigchi-a 1823 | \fi 1824 | } 1825 | \def\@setthanks{\long\def\thanks##1{\par##1\@addpunct.}\thankses} 1826 | \RequirePackage{fancyhdr} 1827 | \if@ACM@review 1828 | \newsavebox{\ACM@linecount@bx} 1829 | \savebox{\ACM@linecount@bx}[4em][t]{\parbox[t]{4em}{% 1830 | \newlength\ACM@linecount@bxht\setlength{\ACM@linecount@bxht}{-\baselineskip} 1831 | \@tempcnta\@ne\relax 1832 | \loop{\color{ACMRed}\scriptsize\the\@tempcnta}\\ 1833 | \advance\@tempcnta by \@ne 1834 | \addtolength{\ACM@linecount@bxht}{\baselineskip} 1835 | \ifdim\ACM@linecount@bxht<\textheight\repeat}} 1836 | \fi 1837 | \def\ACM@linecount{% 1838 | \if@ACM@review 1839 | \begin{picture}(0,0)% 1840 | \put(-26,-22){\usebox{\ACM@linecount@bx}}% 1841 | \end{picture}% 1842 | \fi} 1843 | \def\@shortauthors{\if@ACM@anonymous Anon.\else\shortauthors\fi} 1844 | \def\@headfootfont{% 1845 | \ifcase\ACM@format@nr 1846 | \relax % manuscript 1847 | \sffamily 1848 | \or % acmsmall 1849 | \sffamily 1850 | \or % acmlarge 1851 | \sffamily 1852 | \or % acmtog 1853 | \sffamily 1854 | \or % sigconf 1855 | \sffamily 1856 | \or % siggraph 1857 | \sffamily 1858 | \or % sigplan 1859 | \sffamily 1860 | \or % sigchi 1861 | \sffamily 1862 | \or % sigchi-a 1863 | \sffamily 1864 | \fi} 1865 | \fancypagestyle{standardpagestyle}{% 1866 | \fancyhf{}% 1867 | \renewcommand{\headrulewidth}{\z@}% 1868 | \renewcommand{\footrulewidth}{\z@}% 1869 | \ifcase\ACM@format@nr 1870 | \relax % manuscript 1871 | \fancyhead[LE]{\ACM@linecount\if@ACM@printfolios\thepage\fi}% 1872 | \fancyhead[RO]{\if@ACM@printfolios\thepage\fi}% 1873 | \fancyhead[RE]{\@shortauthors}% 1874 | \fancyhead[LO]{\ACM@linecount\shorttitle}% 1875 | \fancyfoot[RO,LE]{\footnotesize Manuscript submitted to ACM}% 1876 | \or % acmsmall 1877 | \fancyhead[LE]{\ACM@linecount\@headfootfont\@acmArticle\if@ACM@printfolios:\thepage\fi}% 1878 | \fancyhead[RO]{\@headfootfont\@acmArticle\if@ACM@printfolios:\thepage\fi}% 1879 | \fancyhead[RE]{\@headfootfont\@shortauthors}% 1880 | \fancyhead[LO]{\ACM@linecount\@headfootfont\shorttitle}% 1881 | \fancyfoot[RO,LE]{\footnotesize \@journalName, Vol. \@acmVolume, No. 1882 | \@acmNumber, Article \@acmArticle. Publication date: \@acmPubDate.}% 1883 | \or % acmlarge 1884 | \fancyhead[LE]{\ACM@linecount\@headfootfont 1885 | \@acmArticle:\if@ACM@printfolios\thepage\quad\textbullet\quad\fi\@shortauthors}% 1886 | \fancyhead[LO]{\ACM@linecount}% 1887 | \fancyhead[RO]{\@headfootfont 1888 | \shorttitle\quad\textbullet\quad\@acmArticle\if@ACM@printfolios:\thepage\fi}% 1889 | \fancyfoot[RO,LE]{\footnotesize \@journalName, Vol. \@acmVolume, No. 1890 | \@acmNumber, Article \@acmArticle. Publication date: \@acmPubDate.}% 1891 | \or % acmtog 1892 | \fancyhead[LE]{\ACM@linecount\@headfootfont 1893 | \@acmArticle:\if@ACM@printfolios\thepage\quad\textbullet\quad\fi\@shortauthors}% 1894 | \fancyhead[LO]{\ACM@linecount}% 1895 | \fancyhead[RO]{\@headfootfont 1896 | \shorttitle\quad\textbullet\quad\@acmArticle\if@ACM@printfolios:\thepage\fi}% 1897 | \fancyfoot[RO,LE]{\footnotesize \@journalName, Vol. \@acmVolume, No. 1898 | \@acmNumber, Article \@acmArticle. Publication date: \@acmPubDate.}% 1899 | \else % Proceedings 1900 | \fancyfoot[C]{\if@ACM@printfolios\footnotesize\thepage\fi}% 1901 | \fancyhead[LO]{\ACM@linecount\@headfootfont\shorttitle}% 1902 | \fancyhead[RE]{\@headfootfont\@shortauthors}% 1903 | \fancyhead[LE]{\ACM@linecount\@headfootfont\acmConference@shortname, 1904 | \acmConference@date, \acmConference@venue}% 1905 | \fancyhead[RO]{\@headfootfont\acmConference@shortname, 1906 | \acmConference@date, \acmConference@venue}% 1907 | \fi 1908 | \if@ACM@sigchiamode 1909 | \fancyheadoffset[L]{\dimexpr(\marginparsep+\marginparwidth)}% 1910 | \fi 1911 | } 1912 | \pagestyle{standardpagestyle} 1913 | \newdimen\@folio@wd 1914 | \@folio@wd=\z@ 1915 | \newdimen\@folio@ht 1916 | \@folio@ht=\z@ 1917 | \newdimen\@folio@voffset 1918 | \@folio@voffset=\z@ 1919 | \def\@folio@max{1} 1920 | \ifcase\ACM@format@nr 1921 | \relax % manuscript 1922 | \or % acmsmall 1923 | \@folio@wd=45.75pt\relax 1924 | \@folio@ht=1.25in\relax 1925 | \@folio@voffset=.2in\relax 1926 | \def\@folio@max{8} 1927 | \or % acmlarge 1928 | \@folio@wd=43.25pt\relax 1929 | \@folio@ht=79pt\relax 1930 | \@folio@voffset=.55in\relax 1931 | \def\@folio@max{10} 1932 | \fi 1933 | \def\@folioblob{\@tempcnta=\@acmArticleSeq\relax 1934 | \loop 1935 | \ifnum\@tempcnta>\@folio@max\relax 1936 | \advance\@tempcnta by - \@folio@max 1937 | \repeat 1938 | \advance\@tempcnta by -1\relax 1939 | \@tempdima=\@folio@ht\relax 1940 | \multiply\@tempdima by \the\@tempcnta\relax 1941 | \advance\@tempdima by -\@folio@voffset\relax 1942 | \begin{picture}(0,0) 1943 | \makebox[\z@]{\raisebox{-\@tempdima}{% 1944 | \rlap{% 1945 | \raisebox{-0.45\@folio@ht}[\z@][\z@]{% 1946 | \rule{\@folio@wd}{\@folio@ht}}}% 1947 | \parbox{\@folio@wd}{% 1948 | \centering 1949 | \textcolor{white}{\LARGE\bfseries\sffamily\@acmArticle}}}} 1950 | \end{picture}} 1951 | 1952 | \fancypagestyle{firstpagestyle}{% 1953 | \fancyhf{}% 1954 | \renewcommand{\headrulewidth}{\z@}% 1955 | \renewcommand{\footrulewidth}{\z@}% 1956 | \ifcase\ACM@format@nr 1957 | \relax % manuscript 1958 | \fancyhead[L]{\ACM@linecount}% 1959 | \fancyfoot[RO,LE]{\if@ACM@printfolios\small\thepage\fi}% 1960 | \fancyfoot[RE,LO]{\footnotesize Manuscript submitted to ACM}% 1961 | \or % acmsmall 1962 | \fancyfoot[RO,LE]{\footnotesize \@journalName, Vol. \@acmVolume, No. 1963 | \@acmNumber, Article \@acmArticle. Publication date: 1964 | \@acmPubDate.}% 1965 | \fancyhead[LE]{\ACM@linecount\@folioblob}% 1966 | \fancyhead[LO]{\ACM@linecount}% 1967 | \fancyhead[RO]{\@folioblob}% 1968 | \fancyheadoffset[RO,LE]{0.6\@folio@wd}% 1969 | \or % acmlarge 1970 | \fancyfoot[RO,LE]{\footnotesize \@journalName, Vol. \@acmVolume, No. 1971 | \@acmNumber, Article \@acmArticle. Publication date: 1972 | \@acmPubDate.}% 1973 | \fancyhead[RO]{\@folioblob}% 1974 | \fancyhead[LE]{\ACM@linecount\@folioblob}% 1975 | \fancyhead[LO]{\ACM@linecount}% 1976 | \fancyheadoffset[RO,LE]{1.4\@folio@wd}% 1977 | \or % acmtog 1978 | \fancyfoot[RO,LE]{\footnotesize \@journalName, Vol. \@acmVolume, No. 1979 | \@acmNumber, Article \@acmArticle. Publication date: 1980 | \@acmPubDate.}% 1981 | \fancyhead[L]{\ACM@linecount}% 1982 | \else % Conference proceedings 1983 | \fancyhead[L]{\ACM@linecount}% 1984 | \fancyfoot[C]{\if@ACM@printfolios\footnotesize\thepage\fi}% 1985 | \fi 1986 | } 1987 | \renewcommand\section{\@startsection{section}{1}{\z@}% 1988 | {-.75\baselineskip \@plus -2\p@ \@minus -.2\p@}% 1989 | {.25\baselineskip}% 1990 | {\@secfont}} 1991 | \renewcommand\subsection{\@startsection{subsection}{2}{\z@}% 1992 | {-.75\baselineskip \@plus -2\p@ \@minus -.2\p@}% 1993 | {.25\baselineskip}% 1994 | {\@subsecfont}} 1995 | \renewcommand\subsubsection{\@startsection{subsubsection}{3}{10pt}% 1996 | {-.5\baselineskip \@plus -2\p@ \@minus -.2\p@}% 1997 | {-3.5\p@}% 1998 | {\@subsubsecfont\@adddotafter}} 1999 | \renewcommand\paragraph{\@startsection{paragraph}{4}{\parindent}% 2000 | {-.5\baselineskip \@plus -2\p@ \@minus -.2\p@}% 2001 | {-3.5\p@}% 2002 | {\@parfont\@adddotafter}} 2003 | \renewcommand\part{\@startsection{part}{9}{\z@}% 2004 | {-10\p@ \@plus -4\p@ \@minus -2\p@}% 2005 | {4\p@}% 2006 | {\@parfont}} 2007 | \def\section@raggedright{\@rightskip\@flushglue 2008 | \rightskip\@rightskip 2009 | \leftskip\z@skip 2010 | \parindent\z@} 2011 | \def\@secfont{\sffamily\bfseries\section@raggedright\MakeUppercase} 2012 | \def\@subsecfont{\sffamily\bfseries\section@raggedright} 2013 | \def\@subsubsecfont{\sffamily\itshape} 2014 | \def\@parfont{\itshape} 2015 | \setcounter{secnumdepth}{3} 2016 | \ifcase\ACM@format@nr 2017 | \relax % manuscript 2018 | \or % acmsmall 2019 | \or % acmlarge 2020 | \def\@secfont{\sffamily\large\section@raggedright\MakeUppercase} 2021 | \def\@subsecfont{\sffamily\large\section@raggedright} 2022 | \or % acmtog 2023 | \def\@secfont{\sffamily\large\section@raggedright\MakeUppercase} 2024 | \def\@subsecfont{\sffamily\large\section@raggedright} 2025 | \or % sigconf 2026 | \def\@secfont{\bfseries\Large\section@raggedright\MakeUppercase} 2027 | \def\@subsecfont{\bfseries\Large\section@raggedright} 2028 | \or % siggraph 2029 | \def\@secfont{\bfseries\sffamily\Large\section@raggedright\MakeUppercase} 2030 | \def\@subsecfont{\bfseries\sffamily\Large\section@raggedright} 2031 | \or % sigplan 2032 | \def\@secfont{\bfseries\Large\section@raggedright} 2033 | \def\@subsecfont{\bfseries\section@raggedright} 2034 | \renewcommand\subsubsection{\@startsection{subsubsection}{3}{\z@}% 2035 | {-.75\baselineskip \@plus -2\p@ \@minus -.2\p@}% 2036 | {.25\baselineskip}% 2037 | {\@subsubsecfont}} 2038 | \def\@subsubsecfont{\bfseries\section@raggedright} 2039 | \renewcommand\paragraph{\@startsection{paragraph}{4}{\z@}% 2040 | {-.5\baselineskip \@plus -2\p@ \@minus -.2\p@}% 2041 | {-3.5\p@}% 2042 | {\@parfont\@addspaceafter}} 2043 | \def\@parfont{\bfseries\itshape} 2044 | \renewcommand\subparagraph{\@startsection{subparagraph}{5}{\z@}% 2045 | {-.5\baselineskip \@plus -2\p@ \@minus -.2\p@}% 2046 | {-3.5\p@}% 2047 | {\@subparfont\@addspaceafter}} 2048 | \def\@subparfont{\itshape} 2049 | \or % sigchi 2050 | \setcounter{secnumdepth}{1} 2051 | \def\@secfont{\bfseries\sffamily\section@raggedright\MakeUppercase} 2052 | \def\@subsecfont{\bfseries\sffamily\section@raggedright} 2053 | \or % sigchi-a 2054 | \setcounter{secnumdepth}{0} 2055 | \def\@secfont{\bfseries\sffamily\section@raggedright\MakeUppercase} 2056 | \def\@subsecfont{\bfseries\sffamily\section@raggedright} 2057 | \fi 2058 | \def\@adddotafter#1{#1\@addpunct{.}} 2059 | \def\@addspaceafter#1{#1\@addpunct{\enspace}} 2060 | \def\@acmplainbodyfont{\itshape} 2061 | \def\@acmplainindent{\parindent} 2062 | \def\@acmplainheadfont{\scshape} 2063 | \def\@acmplainnotefont{\@empty} 2064 | \ifcase\ACM@format@nr 2065 | \relax % manuscript 2066 | \or % acmsmall 2067 | \or % acmlarge 2068 | \or % acmtog 2069 | \or % sigconf 2070 | \or % siggraph 2071 | \or % sigplan 2072 | \def\@acmplainbodyfont{\itshape} 2073 | \def\@acmplainindent{\z@} 2074 | \def\@acmplainheadfont{\bfseries} 2075 | \def\@acmplainnotefont{\normalfont} 2076 | \or % sigchi 2077 | \or % sigchi-a 2078 | \fi 2079 | \newtheoremstyle{acmplain}% 2080 | {.5\baselineskip\@plus.2\baselineskip 2081 | \@minus.2\baselineskip}% space above 2082 | {.5\baselineskip\@plus.2\baselineskip 2083 | \@minus.2\baselineskip}% space below 2084 | {\@acmplainbodyfont}% body font 2085 | {\@acmplainindent}% indent amount 2086 | {\@acmplainheadfont}% head font 2087 | {.}% punctuation after head 2088 | {.5em}% spacing after head 2089 | {\thmname{#1}\thmnumber{ #2}\thmnote{ {\@acmplainnotefont(#3)}}}% head spec 2090 | \def\@acmdefinitionbodyfont{\normalfont} 2091 | \def\@acmdefinitionindent{\parindent} 2092 | \def\@acmdefinitionheadfont{\itshape} 2093 | \def\@acmdefinitionnotefont{\@empty} 2094 | \ifcase\ACM@format@nr 2095 | \relax % manuscript 2096 | \or % acmsmall 2097 | \or % acmlarge 2098 | \or % acmtog 2099 | \or % sigconf 2100 | \or % siggraph 2101 | \or % sigplan 2102 | \def\@acmdefinitionbodyfont{\normalfont} 2103 | \def\@acmdefinitionindent{\z@} 2104 | \def\@acmdefinitionheadfont{\bfseries} 2105 | \def\@acmdefinitionnotefont{\normalfont} 2106 | \or % sigchi 2107 | \or % sigchi-a 2108 | \fi 2109 | \newtheoremstyle{acmdefinition}% 2110 | {.5\baselineskip\@plus.2\baselineskip 2111 | \@minus.2\baselineskip}% space above 2112 | {.5\baselineskip\@plus.2\baselineskip 2113 | \@minus.2\baselineskip}% space below 2114 | {\@acmdefinitionbodyfont}% body font 2115 | {\@acmdefinitionindent}% indent amount 2116 | {\@acmdefinitionheadfont}% head font 2117 | {.}% punctuation after head 2118 | {.5em}% spacing after head 2119 | {\thmname{#1}\thmnumber{ #2}\thmnote{ {\@acmdefinitionnotefont(#3)}}}% head spec 2120 | \theoremstyle{acmplain} 2121 | \newtheorem{theorem}{Theorem}[section] 2122 | \newtheorem{conjecture}[theorem]{Conjecture} 2123 | \newtheorem{proposition}[theorem]{Proposition} 2124 | \newtheorem{lemma}[theorem]{Lemma} 2125 | \newtheorem{corollary}[theorem]{Corollary} 2126 | \theoremstyle{acmdefinition} 2127 | \newtheorem{example}[theorem]{Example} 2128 | \newtheorem{definition}[theorem]{Definition} 2129 | \theoremstyle{acmplain} 2130 | \def\@proofnamefont{\scshape} 2131 | \def\@proofindent{\indent} 2132 | \ifcase\ACM@format@nr 2133 | \relax % manuscript 2134 | \or % acmsmall 2135 | \or % acmlarge 2136 | \or % acmtog 2137 | \or % sigconf 2138 | \or % siggraph 2139 | \or % sigplan 2140 | \def\@proofnamefont{\itshape} 2141 | \def\@proofindent{\noindent} 2142 | \or % sigchi 2143 | \or % sigchi-a 2144 | \fi 2145 | \renewenvironment{proof}[1][\proofname]{\par 2146 | \pushQED{\qed}% 2147 | \normalfont \topsep6\p@\@plus6\p@\relax 2148 | \trivlist 2149 | \item[\@proofindent\hskip\labelsep 2150 | {\@proofnamefont #1\@addpunct{.}}]\ignorespaces 2151 | }{% 2152 | \popQED\endtrivlist\@endpefalse 2153 | } 2154 | \specialcomment{acks}{% 2155 | \begingroup 2156 | \section*{Acknowledgments} 2157 | \phantomsection\addcontentsline{toc}{section}{Acknowledgments} 2158 | }{% 2159 | \endgroup 2160 | } 2161 | \def\grantsponsor#1#2#3{#2} 2162 | \newcommand\grantnum[3][]{#3% 2163 | \def\@tempa{#1}\ifx\@tempa\@empty\else\space(\url{#1})\fi} 2164 | \if@ACM@screen 2165 | \includecomment{screenonly} 2166 | \excludecomment{printonly} 2167 | \else 2168 | \excludecomment{screenonly} 2169 | \includecomment{printonly} 2170 | \fi 2171 | \if@ACM@anonymous 2172 | \excludecomment{anonsuppress} 2173 | \excludecomment{acks} 2174 | \else 2175 | \includecomment{anonsuppress} 2176 | \fi 2177 | \newcommand\showeprint[2][arxiv]{% 2178 | \def\@tempa{#1}% 2179 | \ifx\@tempa\@empty\def\@tempa{arxiv}\fi 2180 | \def\@tempb{arxiv}% 2181 | \ifx\@tempa\@tempb 2182 | arXiv:\href{http://arxiv.org/abs/#2}{#2}\else arXiv:#2% 2183 | \fi} 2184 | \normalsize\normalfont 2185 | \endinput 2186 | %% 2187 | %% End of file `acmart.cls'. 2188 | -------------------------------------------------------------------------------- /src/core-agda.tex: -------------------------------------------------------------------------------- 1 | \documentclass[acmlarge,fleqn]{acmart}\settopmatter{} 2 | 3 | %% Note: Authors migrating a paper from PACMPL format to traditional 4 | %% SIGPLAN proceedings format should change 'acmlarge' to 5 | %% 'sigplan,10pt'. 6 | 7 | 8 | %% Some recommended packages. 9 | \usepackage{booktabs} %% For formal tables: 10 | %% http://ctan.org/pkg/booktabs 11 | \usepackage{subcaption} %% For complex figures with subfigures/subcaptions 12 | %% http://ctan.org/pkg/subcaption 13 | \usepackage{mathpartir} 14 | 15 | 16 | \makeatletter\if@ACM@journal\makeatother 17 | %% Journal information (used by PACMPL format) 18 | %% Supplied to authors by publisher for camera-ready submission 19 | \acmJournal{PACMPL} 20 | \acmVolume{1} 21 | \acmNumber{1} 22 | \acmArticle{1} 23 | \acmYear{2017} 24 | \acmMonth{1} 25 | \acmDOI{10.1145/nnnnnnn.nnnnnnn} 26 | \startPage{1} 27 | \else\makeatother 28 | %% Conference information (used by SIGPLAN proceedings format) 29 | %% Supplied to authors by publisher for camera-ready submission 30 | \acmConference[PL'18]{ACM SIGPLAN Conference on Programming Languages}{January 01--03, 2018}{New York, NY, USA} 31 | \acmYear{2018} 32 | \acmISBN{978-x-xxxx-xxxx-x/YY/MM} 33 | \acmDOI{10.1145/nnnnnnn.nnnnnnn} 34 | \startPage{1} 35 | \fi 36 | 37 | 38 | %% Copyright information 39 | %% Supplied to authors (based on authors' rights management selection; 40 | %% see authors.acm.org) by publisher for camera-ready submission 41 | \setcopyright{none} %% For review submission 42 | %\setcopyright{acmcopyright} 43 | %\setcopyright{acmlicensed} 44 | %\setcopyright{rightsretained} 45 | %\copyrightyear{2017} %% If different from \acmYear 46 | 47 | 48 | %% Bibliography style 49 | \bibliographystyle{ACM-Reference-Format} 50 | %% Citation style 51 | %% Note: author/year citations are required for papers published as an 52 | %% issue of PACMPL. 53 | \citestyle{acmauthoryear} %% For author/year citations 54 | 55 | \RequirePackage{ifthen} 56 | \RequirePackage{amssymb} 57 | \RequirePackage{amsfonts} 58 | \RequirePackage{stmaryrd} % \shortuparrow 59 | 60 | 61 | \input{macros} 62 | 63 | % COVERED by option fleqn: 64 | % 65 | % controlling flushleft/center for math displays 66 | % http://www.golatex.de/wechsel-zwischen-leqno-und-reqno-fleqn-uvm-t2488.html 67 | \makeatletter 68 | %\def\leqn{\tagsleft@true} 69 | %\def\reqn{\tagsleft@false} 70 | \def\fleq{\@fleqntrue\let\mathindent\@mathmargin \@mathmargin=\leftmargini} 71 | \def\cneq{\@fleqnfalse} 72 | \g@addto@macro{\endsubequations}{\addtocounter{equation}{-1}} 73 | \makeatother 74 | 75 | \begin{document} 76 | 77 | %% Title information 78 | \title[Core Agda]{Specification of Core Agda} %% [Short Title] is optional; 79 | %% when present, will be used in 80 | %% header instead of Full Title. 81 | \titlenote{with title note} %% \titlenote is optional; 82 | %% can be repeated if necessary; 83 | %% contents suppressed with 'anonymous' 84 | \subtitle{Subtitle} %% \subtitle is optional 85 | \subtitlenote{with subtitle note} %% \subtitlenote is optional; 86 | %% can be repeated if necessary; 87 | %% contents suppressed with 'anonymous' 88 | 89 | 90 | %% Author information 91 | %% Contents and number of authors suppressed with 'anonymous'. 92 | %% Each author should be introduced by \author, followed by 93 | %% \authornote (optional), \orcid (optional), \affiliation, and 94 | %% \email. 95 | %% An author may have multiple affiliations and/or emails; repeat the 96 | %% appropriate command. 97 | %% Many elements are not rendered, but should be provided for metadata 98 | %% extraction tools. 99 | 100 | %% Author with single affiliation. 101 | \author{The Agda Specification Team} 102 | \authornote{with author1 note} %% \authornote is optional; 103 | %% can be repeated if necessary 104 | \orcid{nnnn-nnnn-nnnn-nnnn} %% \orcid is optional 105 | \affiliation{ 106 | \position{Position1} 107 | \department{Department1} %% \department is recommended 108 | \institution{Institution1} %% \institution is required 109 | \streetaddress{Street1 Address1} 110 | \city{City1} 111 | \state{State1} 112 | \postcode{Post-Code1} 113 | \country{Country1} 114 | } 115 | \email{first1.last1@inst1.edu} %% \email is recommended 116 | 117 | %% Author with two affiliations and emails. 118 | \author{First2 Last2} 119 | \authornote{with author2 note} %% \authornote is optional; 120 | %% can be repeated if necessary 121 | \orcid{nnnn-nnnn-nnnn-nnnn} %% \orcid is optional 122 | \affiliation{ 123 | \position{Position2a} 124 | \department{Department2a} %% \department is recommended 125 | \institution{Institution2a} %% \institution is required 126 | \streetaddress{Street2a Address2a} 127 | \city{City2a} 128 | \state{State2a} 129 | \postcode{Post-Code2a} 130 | \country{Country2a} 131 | } 132 | \email{first2.last2@inst2a.com} %% \email is recommended 133 | \affiliation{ 134 | \position{Position2b} 135 | \department{Department2b} %% \department is recommended 136 | \institution{Institution2b} %% \institution is required 137 | \streetaddress{Street3b Address2b} 138 | \city{City2b} 139 | \state{State2b} 140 | \postcode{Post-Code2b} 141 | \country{Country2b} 142 | } 143 | \email{first2.last2@inst2b.org} %% \email is recommended 144 | 145 | 146 | %% Paper note 147 | %% The \thanks command may be used to create a "paper note" --- 148 | %% similar to a title note or an author note, but not explicitly 149 | %% associated with a particular element. It will appear immediately 150 | %% above the permission/copyright statement. 151 | \thanks{with paper note} %% \thanks is optional 152 | %% can be repeated if necesary 153 | %% contents suppressed with 'anonymous' 154 | 155 | 156 | %% Abstract 157 | %% Note: \begin{abstract}...\end{abstract} environment must come 158 | %% before \maketitle command 159 | \begin{abstract} 160 | This document specifies the abstract syntax, evaluation rules, and 161 | typing rules of Core Agda, the basic type theory underlying Agda. 162 | \end{abstract} 163 | 164 | 165 | %% 2012 ACM Computing Classification System (CSS) concepts 166 | %% Generate at 'http://dl.acm.org/ccs/ccs.cfm'. 167 | \begin{CCSXML} 168 | 169 | 170 | 10011007.10011006.10011008 171 | Software and its engineering~General programming languages 172 | 500 173 | 174 | 175 | 10003456.10003457.10003521.10003525 176 | Social and professional topics~History of programming languages 177 | 300 178 | 179 | 180 | \end{CCSXML} 181 | 182 | \ccsdesc[500]{Software and its engineering~General programming languages} 183 | \ccsdesc[300]{Social and professional topics~History of programming languages} 184 | %% End of generated code 185 | 186 | 187 | %% Keywords 188 | %% comma separated list 189 | \keywords{Agda, dependent types, specification} %% \keywords is optional 190 | 191 | 192 | %% \maketitle 193 | %% Note: \maketitle command must come after title commands, author 194 | %% commands, abstract environment, Computing Classification System 195 | %% environment and commands, and keywords command. 196 | \maketitle 197 | 198 | 199 | \section{Introduction} 200 | 201 | Agda 2 has been developed since 2005, and been released 2007. So far, 202 | no specification has been given. This document attempts to specify 203 | the core components of Agda. 204 | 205 | 206 | \section{Syntax} 207 | \label{sec:syntax} 208 | 209 | \subsection{Terms} 210 | 211 | We distinguish global names $f,F,D,c,R,\pi$ bound in the signature $\Sigma$ from 212 | local variables $x,y,z$ bound in typing contexts $\Gamma$. Local 213 | variables are represented by de Bruijn indices in the implementation, 214 | but for the sake of readability we stick to a named presentation here. 215 | We silently rename bound variables to avoid clashes with free variables 216 | (Barendregt convention). We write $x \# E$ (``$x$ fresh for $E$'') 217 | to express that $x$ is a 218 | fresh variable with regard to syntactic entity $E$. 219 | 220 | For orientation of the reader, we use different letters to represent 221 | different purposes of global names. However, they share the same name 222 | space and are not distinguished syntactically. 223 | \[ 224 | \begin{array}{l@{\hspace{2cm}}l} 225 | D & \text{data type name} \\ 226 | R & \text{record type name} \\ 227 | F & \text{data type or record type name} \\ 228 | f & \text{function name} \\ 229 | \pi & \text{projection name (overloaded)} \\ 230 | c & \text{data/record constructor name (overloaded)} \\ 231 | \end{array} 232 | \] 233 | Projection and constructor names can be overloaded and are resolved by 234 | the type checker. We write $R.\pi$ for a resolved record projection 235 | name, $R.c$ for a resolved record constructor name and $D.c$ for a 236 | resolved data constructor name. 237 | 238 | \emph{Terms} use a spine form for eliminations and thus are kept $\beta$ and 239 | projection normal form. This means that terms cannot be a 240 | $\beta$-redex $(\lambda x.\,v)\,u$ nor a projection of a record value 241 | $\cpi\,\vec v\,.\pi$. 242 | \[ 243 | \begin{array}{lrl@{\hspace{2cm}}l} 244 | t, u, v 245 | &::=& x\; \bar{e} & \text{variables (eliminated by $\bar{e}$)} \\ 246 | & | & f\; \bar{e} & \text{defined function (eliminated by $\bar{e}$), includes postulates} \\ 247 | & | & c\; \bar{v} & \text{data constructor applied to arguments $\bar{v}$} \\ 248 | & | & \cpi\; \bar{v} & \text{record constructor with fields $\vec\pi$ applied to arguments $\bar{v}$} \\ 249 | & | & \lambda x.\, v & \text{lambda abstraction} \\ 250 | & | & F\; \bar{v} & \text{data/record type name applied to arguments $\bar v$} \\ 251 | & | & \funT x U V & \text{dependent function type} \\ 252 | & | & s & \text{Sort} \\ 253 | \end{array} 254 | \] 255 | We use uppercase letters $T,U,V$ for terms that stand for types. 256 | 257 | \emph{Eliminations} $e$ exists for functions and records. 258 | Functions are eliminated by application, records by projection. 259 | \[ 260 | \begin{array}{lrl@{\hspace{2cm}}l} 261 | e &::=& @u & \text{application to term $u$} \\ 262 | & | & .\pi & \text{taking projection $\pi$} \\ 263 | \end{array} 264 | \] 265 | Other forms of elimination, like definiting a function by cases over 266 | some data structure, need pattern matching, which is supported in 267 | function declarations, but not in terms directly. 268 | 269 | \emph{Sorts} $s$ are the types of types. 270 | Here, we model a predicative hierarchy of universes $\Set_0 : \Set_1 : \Set_2 : \dots$. 271 | \[ 272 | \begin{array}{lrl@{\hspace{2cm}}l} 273 | s & ::= & \Set_n & \text{universe of types of level $n \in \NN$} \\ 274 | \end{array} 275 | \] 276 | We define \emph{sort supremum} \fbox{$s \sqcup s'$} by $\Set_i \sqcup \Set_j = \Set_{\max(i,j)}$. 277 | 278 | 279 | \subsection{Telescopes and patterns} 280 | 281 | \emph{Telescopes} $\Delta$ are like typing contexts $\Gamma$ (see below), but right-associative. 282 | Telescopes are used \eg{} for parameter lists in declarations. 283 | \[ 284 | \begin{array}{lrl@{\hspace{2cm}}l} 285 | \Delta &::=& \cempty & \text{empty telescope} \\ 286 | & |& (x : T)\Delta & \text{non-empty telescope; $\Delta$ can depend on $x$} \\ 287 | \end{array} 288 | \] 289 | Notation: iterated function types \fbox{$\Delta \to T$}, defined by recursion on $\Delta$: 290 | \[ 291 | \begin{array}{rrl@{\hspace{2cm}}l} 292 | \cempty \to T &=& T &\\ 293 | (x : U)\Delta \to T &=& (x : U) \to (\Delta \to T) & \\ 294 | \end{array} 295 | \] 296 | 297 | \emph{Patterns} $p$ are used for definition by cases. 298 | Patterns are made from variables $x$ and data constructors $c$, 299 | but can also contain arbitrary terms $u$ that do not bind any variables. 300 | The latter form is called \emph{inaccessible} pattern, or, due to the 301 | concrete syntax used in Agda, \emph{dot} pattern. 302 | \[ 303 | \begin{array}{lrl@{\hspace{2cm}}l} 304 | p &::=& x & \text{variable pattern} \\ 305 | & | & c\; \bar{p} & \text{constructor pattern} \\ 306 | & | & \dotp u & \text{inaccessible pattern (aka dot pattern)} \\ 307 | \end{array} 308 | \] 309 | 310 | \emph{Copatterns} are to eliminations what patterns are to arguments. 311 | We match eliminations against copatterns like we match arguments against patterns. 312 | \[ 313 | \begin{array}{lrl@{\hspace{2cm}}l} 314 | q &::= & @p & \text{application pattern} \\ 315 | & | & .\pi & \text{projection pattern}\\ 316 | \end{array} 317 | \] 318 | 319 | \subsection{Declarations} 320 | \label{sec:declarations} 321 | 322 | Agda has three main \emph{declaration} forms that introduce global names into 323 | the signature $\Sigma$. Function declarations introduce global 324 | functions $f$ defined by pattern matching. A function declaration 325 | consists of a type signature $\vts$ and a list of function clauses 326 | $\vec\vcl$. The signature has to preceed the clauses, but between 327 | them other declarations are allowed, in order to facilitate mutually 328 | recursive definitions. Data (type) declarations consist of a data 329 | signature $\vds$ and a data definition $\vdd$, which have to appear in 330 | this order but can also be appart, to realize inductive-recursive 331 | definitions, for instance. Similarly record (type) declarations 332 | consist of a record signature $\vrs$ and a record declaration $\vrd$. 333 | \[ 334 | \begin{array}{lrl@{\hspace{2cm}}l} 335 | d & ::= & \mathit{ts} & \mbox{type signature} 336 | \\ & \mid & \mathit{cl} & \mbox{function clause} 337 | \\ & \mid & \mathit{ds} & \mbox{data signature} 338 | \\ & \mid & \mathit{dd} & \mbox{data definition} 339 | \\ & \mid & \mathit{rs} & \mbox{record signature} 340 | \\ & \mid & \mathit{rd} & \mbox{record definition}. 341 | \end{array} 342 | \] 343 | 344 | \emph{Type signatures} consist of a global name $f$ for the function and its type $T$. 345 | \[\begin{array}{lrll} 346 | ts &::=& f : T &\\ 347 | \end{array} \] 348 | A \emph{function clause} for $f$ consists of a copattern spine $\vec q$ and a right hand side $\vrhs$. 349 | The pattern variables of $\vec q$ are bound in context $\Gamma$. 350 | The right hand side is either empty or has evidence that the case is impossible. 351 | \[ 352 | \begin{array}{lrll} 353 | cl &::=& \Gamma \rhd f\; \bar{q} : U \; = \vrhs &\\ 354 | \vrhs & ::= & t & \text{term (clause body)}\\ 355 | & \mid& \absurd \vec x & (\vec x \not=\emptyset) 356 | \text{ absurd clause (each variable $x_i$ has empty type)} \\ 357 | \end{array} 358 | \] 359 | 360 | A \emph{data signature} consists of a name $D$ for the data type, 361 | a list of its parameters $\Delta$, 362 | a list of itsindices $\Delta'$, 363 | and its target sort $s$. 364 | \[ 365 | \textbf{data}\; D\; \Delta : \Delta' \to s 366 | \] 367 | \emph{Data definitions} repeat the parameter telescope $\Delta$, 368 | since parameters will be mentioned in the types $T$ of the constructors $c$. 369 | \[ 370 | \textbf{data}\; D\; \Delta \;\textbf{where}\; \overline{c : T} 371 | \] 372 | 373 | \emph{Record signature}s carry, in contrast to data signatures, 374 | no index telescope $\Delta'$, as Agda does not support indexed record types. 375 | \[ 376 | \begin{array}{@{}l} 377 | \textbf{record}\; R\; \Delta : s 378 | \end{array} 379 | \] 380 | \emph{Record declarations} supply a record constructor name $c$ 381 | and a list of fields $\pi$ with their types $T$. 382 | \[ 383 | \begin{array}{@{}l} 384 | \textbf{record}\; R\; \Delta \; \textbf{where} \\ 385 | \qquad \textbf{constructor}\; c \\ 386 | \qquad \textbf{field}\; \overline{\pi : T} \\ 387 | \end{array} 388 | \] 389 | 390 | Declaration checking moves declarations to the signature $\Sigma$, 391 | but not literally, but in a refined form. 392 | A signature is a list of signature declarations (used internally). 393 | \[ 394 | \begin{array}{lcl} 395 | \Sigma & ::= & \overline{d_\Sigma} \\ 396 | d_\Sigma & ::= & \datasig \\ 397 | & \mid & \recsig \\ 398 | & \mid & \funsig 399 | \end{array} 400 | \] 401 | TODO: UPDATE SIGNATURE DECLARATIONS. 402 | 403 | 404 | 405 | 406 | \section{Typing rules} 407 | \label{sec:typing} 408 | 409 | 410 | \subsection{Auxiliary judgements} 411 | 412 | Hereditary substitution \fbox{$u[\sigma] = v$} 413 | is defined simultaneously with active elimination (see below). 414 | \begin{gather*} 415 | \inferrule 416 | {\sigma(x) \bang \overline{e}[\sigma] = v} 417 | {(x \overline{e})[\sigma] = v} 418 | \qquad 419 | \inferrule 420 | {\overline{e}[\sigma] = \overline{e}'} 421 | {(f \overline{e})[\sigma] = f \overline{e}'} 422 | \qquad 423 | \inferrule 424 | {\overline{v}[\sigma] = \overline{v}'} 425 | {(h \overline{v})[\sigma] = h \overline{v}'} 426 | \ h ::= F \mid c \mid \cpi 427 | \qquad 428 | \rux{v[\sigma] = v' 429 | }{(\lam x v)[\sigma] = \lam x v' 430 | }{x \# \sigma} 431 | \\ 432 | \rux{U[\sigma] = U' \qquad 433 | V[\sigma] = V' 434 | }{(\piT x U V)[\sigma] = \piT x {U'}{V'} 435 | }{x \# \sigma} 436 | \qquad 437 | \ru{}{s[\sigma] = s} 438 | \end{gather*} 439 | Active elimination \fbox{$t \bang \overline{e} = v$}\ . 440 | \begin{gather*} 441 | \ru 442 | { v[u/x] \bang \overline{e} = v'} 443 | {\lambda x.v \bang @u \overline{e} = v'} 444 | \qquad 445 | \inferrule 446 | {v_i \bang \overline{e} = v'} 447 | {\cpi \overline{v} \bang .\pi_i \overline{e} = v'} 448 | \qquad 449 | \ru{}{t \bang \cdot = t} 450 | \qquad 451 | \ru{}{x \overline{e} \bang \overline{e}' = x \overline{e} 452 | \overline{e}'} 453 | \qquad 454 | \ru{}{f \overline{e} \bang \overline{e}' = f \overline{e} \overline{e}'} 455 | %% We let constructors, record and data types be fully applied. 456 | % \qquad 457 | % \ru{}{z \overline{v} \bang @ \overline{u} = z \overline{v} 458 | % \overline{u}}\ z ::= c \mid F 459 | % \qquad 460 | % \ru{}{D \overline{v} \bang @ \overline{u} = D \overline{v} 461 | % \overline{u}} 462 | % \qquad 463 | % \ru{}{R \overline{v} \bang @ \overline{u} = R \overline{v} \overline{u}} 464 | \end{gather*} 465 | Note that since constructor and data and record type terms are fully 466 | applied, we do not need to give rules for adding further applications 467 | to these. Of course, function type and sort expressions cannot be applied either. 468 | % \[ 469 | % \[ 470 | % \begin{array}{ll} 471 | % & 472 | % \\ 473 | % & 474 | % \\ 475 | % \inferrule{c \overline{v} u \bang \overline{e} = v'} 476 | % {\overline{v} \bang @u \overline{e} = v'} 477 | % & 478 | % \\ 479 | % & 480 | % \text{ everything else undefined } 481 | % \end{array} 482 | % \] 483 | 484 | 485 | Function type application \fbox{$T \twobang \vec u = U$}. 486 | \[ 487 | \ru{}{T \twobang . = T} 488 | \qquad 489 | \ru{\subst u x V \twobang \vec u = T 490 | }{\funT x U V \twobang u \vec u = T} 491 | \] 492 | 493 | % Type elimination \fbox{$\Sigma \der (t : T) \twobang e = U$}. 494 | % \[ 495 | % \ru{}{\Sigma \der (t : \funT x U V) \twobang @u = \subst u x V} 496 | % \qquad 497 | % \ru{\Sigma \der R.\pi : U 498 | % }{} 499 | % \] 500 | 501 | Signature lookup \fbox{$\Sigma \der z : T$}. 502 | \begin{gather*} 503 | \ru{\funsig \in \Sigma 504 | }{\Sigma \der f : T} 505 | \\ 506 | \ru{\datasig \in \Sigma 507 | }{\Sigma \der D : \Delta \to \Delta' \to s} 508 | \qquad 509 | \ru{\datasig \in \Sigma 510 | }{\Sigma \der D.c_i : \Delta \to T_i} 511 | \\ 512 | \ru{\recsig \in \Sigma 513 | }{\Sigma \der R : \Delta \to s} 514 | \\ 515 | \ru{\recsig \in \Sigma 516 | }{\Sigma \der R.c : \Delta \to T} 517 | \qquad 518 | \ru{\recsig \in \Sigma 519 | }{\Sigma \der R.\pi_i : \Delta \to T_i} 520 | \end{gather*} 521 | 522 | 523 | 524 | 525 | \subsection{Typing judgements for expressions} 526 | 527 | 528 | Typing contexts. 529 | \[ 530 | \begin{array}{lcl} 531 | \Gamma & ::= & \cdot \mid \Gamma, x:T \\ 532 | \Delta & ::= & \cdot \mid x:T, \Delta 533 | \end{array} 534 | \] 535 | 536 | Well-typed contexts \fbox{$\ders \Gamma$}. 537 | \[ 538 | % Contexts 539 | \ru{}{\ders \cdot} 540 | \qquad 541 | \inferrule 542 | {\ders \Gamma \qquad \Gamma \ders T:s} 543 | {\ders \Gamma, x : T} \ x \not\in \dom(\Gamma) 544 | \] 545 | 546 | Well-typed telescopes \fbox{$\Gamma \ders \Delta$}. 547 | \[ 548 | % Telescopes 549 | \inferrule 550 | {\ders \Gamma} 551 | {\Gamma \ders \cdot} 552 | \qquad 553 | \inferrule{\Gamma \ders T : s \qquad \Gamma, x:T \ders \Delta} 554 | {\Gamma \ders (x:T) \Delta} 555 | % \qquad x \in \dom{\Delta} \\ 556 | \] 557 | 558 | Spine typing \fbox{$\Gamma \mid u : U \ders \overline{e} : T$} 559 | ($u$ is neutral). 560 | \begin{gather*} 561 | % Check spine 562 | \ru{} 563 | {\Gamma \mid u:U \ders \cdot : U} 564 | \\ 565 | \inferrule 566 | {\Gamma \ders u:U \qquad \Gamma \mid t@u:T[u/x] \ders \overline{e} : V} 567 | {\Gamma \mid t : \Pi x:U.T \ders @u \overline{e}:V} 568 | \qquad 569 | \ru{\Sigma \der R.\pi = T \qquad 570 | T \twobang (\vec u,t) = U \qquad 571 | \Gamma \mid t.\pi : U \ders \overline{e} : V 572 | }{\Gamma \mid t : R\overline{u} \ders .\pi \overline{e} : V} 573 | % \inferrule{ \Gamma \vdash (t:R\overline{u}).\pi = T' 574 | % \qquad \Gamma \mid t.\pi :T' \ders \overline{e}:V} 575 | % { \Gamma \mid t : R \overline{u} \ders .\pi \overline{e}:V} 576 | % \\ 577 | % \Sigma \vdash (t:R\overline{u}).\pi = T 578 | % \\ 579 | % (t: \Sigma x:A.B).\proj_2 = B[t.\proj_1/x] & 580 | % \\ 581 | % \inferrule{T[u/x] \twobang \overline{e} = v} 582 | % {(\Pi x:U.T) \twobang @ u \overline{e} = v} 583 | \\ 584 | \ru{\Gamma \mid u : U \ders \vec e : V \qquad 585 | \Gamma \ders U = U' : s 586 | }{\Gamma \mid u : U' \ders \vec e : V} 587 | \end{gather*} 588 | 589 | 590 | Term typing \fbox{$\Gamma \ders t : T$} 591 | \begin{gather*} 592 | \inferrule 593 | {(x:U)\in \Gamma \qquad \Gamma \mid x:U \ders \overline{e}:T} 594 | {\Gamma \ders x \overline{e} : T} 595 | \qquad 596 | \ru{\Sigma \vdash f : T \qquad 597 | \Gamma \mid f : T \ders \vec e : V 598 | }{\Gamma \ders f \vec e : V} 599 | \\ 600 | \ru{\Sigma \der D.c : T \qquad 601 | T \twobang \vec u = T' \qquad 602 | \Gamma \mid c : T' \ders \vec v : D \, \vec u \, \vec t 603 | }{\Gamma \ders c\,\vec v : D\,\vec u\,\vec t} 604 | \qquad 605 | \ru{\Sigma \vdash z : T \qquad 606 | \Gamma \mid z : T \ders \vec u : V 607 | }{\Gamma \ders z \, \vec u : V} \ z ::= D \mid R 608 | \\ 609 | \ru{\Gamma, x \of U \ders v : V 610 | }{\Gamma \ders \lambda x.v : \funT x U V} 611 | \qquad 612 | \ru{\Gamma \ders U : s \qquad 613 | \Gamma, x \of U \ders V : s' 614 | }{\Gamma \ders \piT x U V : s \sqcup s'} 615 | \end{gather*} 616 | 617 | 618 | Conversion 619 | \fbox{$\Gamma \ders t = t' : T$}\ . 620 | \begin{gather*} 621 | % 622 | \ru{f\,\overline q = t \in \Sigma \qquad 623 | (\overline{\embp q})[\sigma] = \overline e \qquad 624 | t[\sigma] = v \qquad 625 | \Gamma \ders f\,\overline e : T \qquad 626 | \Gamma \ders v : T 627 | }{\Gamma \ders f \overline e = v : T} 628 | % \ru{\Sigma \der t \red t' \qquad 629 | % \Gamma \ders t : T \qquad 630 | % \Gamma \ders t' : T 631 | % }{\Gamma \ders t = t' : T} 632 | \\ 633 | \ru{\Gamma \ders t : \funT x U V 634 | }{\Gamma \ders t = \lambda x. t\, x : \funT x U V} 635 | \qquad 636 | \ru{\Gamma \ders t : R\, \vec u \qquad 637 | \Sigma \der R.\cpi : T 638 | }{\Gamma \ders t = \cpi\,\overline{t .\pi} : R\,\vec u} 639 | \\ 640 | \end{gather*} 641 | And many boring rules (equivalence, congruence) 642 | and rules for elimination equality 643 | \fbox{$\Gamma \mid t : T \ders \vec e = \vec e' : V$} 644 | 645 | \newpage 646 | 647 | \subsection{Type emptiness} 648 | 649 | \fbox{$\Gamma \der t \apart t' : T$} In context $\Gamma$, terms $t$ and $t'$ of type $T$ cannot ever be unified. 650 | \begin{gather*} 651 | \rux{\Gamma \der c \vec v : T \qquad \Gamma \der c' \vec v' : T 652 | }{\Gamma \der c \vec v \apart c' \vec v' : T 653 | }{c \not= c'} 654 | \qquad 655 | \ru{\Gamma \der t \apart t' : T 656 | }{\Gamma \der t' \apart t : T} 657 | \qquad 658 | \ru{\Gamma \der t \apart t' : T \qquad \Gamma \der t' = t'' : T 659 | }{\Gamma \der t \apart t'' : T} 660 | \\ 661 | \ru{\Sigma \der c : \Delta \to \Delta' \to D\,\Delta \,\vec u' \qquad 662 | % \subst{\vec t}\Delta{\vec u'} = \vec u'' \qquad 663 | \subst{\vec t}\Delta{\Delta'} = \Delta'' \qquad 664 | \Gamma \der \vec v \apart \vec v' : \Delta'' 665 | }{\Gamma \der c \vec v \apart c \vec v' : D\,\vec t\,\vec u} 666 | % \ru{\Sigma \der c : \Delta \to \Delta' \to D\,\Delta \,\vec u' \qquad 667 | % \subst{\vec t}\Delta{\vec u'} = \vec u'' \qquad 668 | % \subst{\vec t}\Delta{\Delta'} = \Delta'' \qquad 669 | % \Gamma \mid c : \Delta' \to D\,\vec t\,\vec u'' \der \vec v \apart \vec v' : D\,\vec t\,\vec u'' 670 | % }{\Gamma \der c \vec v \apart c \vec v' : D\,\vec t\,\vec u} 671 | \qquad 672 | \ru{\Gamma \der t \apart t' : T \qquad \Gamma \der T = T' : s 673 | }{\Gamma \der t \apart t' : T'} 674 | \end{gather*} 675 | \fbox{$\Gamma \der \vec v \apart \vec v' : \Delta$} 676 | In context $\Gamma$, argument lists $\vec v$ and $\vec v'$ of telescope $\Delta$ cannot ever be unified. 677 | \begin{gather*} 678 | \ru{\Gamma \der v \apart v' : T 679 | }{\Gamma \der v\,\vec v \apart v'\,\vec v' : (x \of T)\Delta} 680 | \\[2ex] 681 | \ru{\Gamma \der v = v' : T \qquad 682 | \subst v x \Delta = \Delta' \qquad 683 | \Gamma \der \vec v \apart \vec v' : \Delta' 684 | }{\Gamma \der v \vec v \apart v' \vec v' : (x \of T) \Delta} 685 | \qquad 686 | \ru{%\Gamma \der \Delta \qquad 687 | x \not\in\FV(\Delta) \qquad 688 | % \subst {v_1} x \Delta = \Delta_1 \qquad 689 | % \subst {v_2} x \Delta = \Delta_2 \qquad 690 | % \Gamma \der \Delta_1 = \Delta_2 \qquad 691 | \Gamma \der \vec v \apart \vec v' : \Delta 692 | }{\Gamma \der v \vec v \apart v' \vec v' : (x \of T) \Delta} 693 | \end{gather*} 694 | 695 | \fbox{$\Gamma \der T \noinstance U$} 696 | \begin{gather*} 697 | \ru{\subst {\vec u} \Delta {\Delta'} = \Delta'' \qquad 698 | \subst {\vec u} \Delta {\vec v'} = \vec v'' \qquad 699 | \Gamma, \Delta'' \der \vec v'' \apart \vec v : \Delta'' 700 | }{\Gamma \der \Delta \to \Delta' \to D\,\Delta\,\vec v' \noinstance D\,\vec u\,\vec v 701 | } 702 | \qquad 703 | \end{gather*} 704 | 705 | \fbox{$\Gamma \der T \mempty$}: In context $\Gamma$, data type $T$ has no canonical inhabitants. 706 | \begin{gather*} 707 | \ru{\forall c.\ \Sigma \der D.c : U \Rightarrow \Gamma \der U \noinstance T 708 | }{\Gamma \der T \mempty 709 | } 710 | \qquad 711 | \ru{\Gamma \der T \mempty \qquad \Gamma \der T = T' : s 712 | }{\Gamma \der T' \mempty} 713 | \end{gather*} 714 | 715 | \begin{verbatim} 716 | 717 | \end{verbatim} 718 | \subsection{Declaration typing} 719 | 720 | (Co)pattern variables \fbox{$PV(p) = \bar{x}$} 721 | \begin{gather*} 722 | PV(x) = \{ x \} \qquad PV (\lceil v \rceil) = \emptyset \qquad PV(c\; \bar{p}) = PV(\bar{p}) \\ 723 | PV(\bar{p}) = \biguplus_i PV(p_i) \\ 724 | PV(.\pi) = \emptyset \qquad PV(@p) = PV(p) \\ 725 | PV(\bar{q}) = \biguplus_i PV(q_i) \\ 726 | \end{gather*} 727 | 728 | 729 | Embedding of (co)patterns into terms \fbox{$\lceil p \rceil = t$} 730 | \begin{gather*} 731 | \lceil x \rceil = x \qquad 732 | \lceil \lfloor v \rfloor \rceil = v \qquad 733 | \lceil c\; \bar{p} \rceil = c\; \lceil \bar{p} \rceil \\ 734 | \lceil .\pi \rceil = .\pi \qquad 735 | \lceil @p \rceil = @ \lceil p \rceil 736 | \end{gather*} 737 | 738 | 739 | %Pattern typing \fbox{$\Gamma \ders p : U \leadsto \Gamma'$} 740 | %\begin{gather*} 741 | %\ru{\Gamma \ders U : s} 742 | % {\Gamma \ders x : U \leadsto (x : U)} x \not\in dom(\Gamma) \\ 743 | %\ru{\Sigma \der D.c : T \qquad T\twobang\bar{u} = T' \qquad \Gamma \mid c : T \ders @\bar{p} : D\; \bar{u}\; \bar{v} \leadsto \Gamma'} 744 | % {\Gamma \ders c\; \bar{p} : D\; \bar{u}\; \bar{v} \leadsto \Gamma'} 745 | %\end{gather*} 746 | % 747 | % 748 | %Copattern typing \fbox{$\Gamma \mid f : T \ders \bar{q} : U \leadsto \Gamma'$} 749 | %\begin{gather*} 750 | %\ru{\Gamma \vdash p : U \leadsto \Gamma_1 \qquad \Gamma,\Gamma' | h\; @p : V[p/x] \ders \bar{q} : T \leadsto \Gamma_2} 751 | % {\Gamma \mid h : \piT x U V \ders @p\; \bar{q} : T \leadsto \Gamma_1 , \Gamma_2} 752 | %\end{gather*} 753 | 754 | Clause typing \fbox{$f : T \ders cl$} 755 | \begin{gather*} 756 | \ru{PV(\bar{q}) = dom(\Gamma) \qquad 757 | \Gamma \mid f : T \ders \lceil \bar{q} \rceil : U \qquad 758 | \Gamma \ders t : U} 759 | {f : T \ders (\Gamma \rhd f\; \bar{q} : U = t)} 760 | \\ 761 | \ru{PV(\bar{q}) = dom(\Gamma) \qquad 762 | \Gamma \mid f : T \ders \lceil \bar{q} \rceil : U \qquad 763 | \forall i.\ \Gamma \ders \Gamma(x_i) \mempty} 764 | {f : T \ders (\Gamma \rhd f\; \bar{q} : U = \absurd \vec x)} 765 | \end{gather*} 766 | 767 | Constructor typing \fbox{$\Delta \mid U : \Delta' \to s \ders c : T$} 768 | \begin{gather*} 769 | \ru{\Delta \ders \Gamma \to T : s \qquad 770 | \Delta,\Gamma \ders \bar{v} : \Delta' \qquad 771 | \Delta,\Gamma \ders T = U\; \bar{v} : s} 772 | {\Delta \mid U : \Delta' \to s \ders c : \Gamma \to T} 773 | \end{gather*} 774 | 775 | % TODO: add syntax for lone declarations without a definition (i.e. with no 'where'). 776 | Declaration typing \fbox{$\Sigma \der d \leadsto \Sigma'$} 777 | \begin{gather*} 778 | \ru{f \not\in \Sigma \qquad \cdot \ders T : s} 779 | {\Sigma \der f : T \leadsto \Sigma, function\; f : T\; where\; \cdot} 780 | \\ 781 | \ru{\forall i.\; f : T \ders cl_i} 782 | {\Sigma[function\; f : T\; where\; \cdot] \der \overline{cl} \leadsto \Sigma[function\; f : T\; where\; \overline{cl}]} 783 | \\ 784 | \ru{D \not\in \Sigma \qquad \ders \Delta \qquad \Delta \ders \Delta'} 785 | {\Sigma \der data\; D\; \Delta : \Delta' \to s \leadsto \Sigma, data\; D\; \Delta : \Delta' \to s} 786 | \\ 787 | \ru{\forall i.\; \Delta \mid D\; \Delta : \Delta' \to s \vdash c_i : T_i} 788 | {\Sigma[data\; D; \Delta : \Delta' \to s] \ders data\; D\; \Delta\; where\; \overline{c : T} \leadsto \Sigma[data\; D\; \Delta : \Delta' \to s\; where\; \overline{c : T}]} 789 | \\ 790 | \ru{R \not\in \Sigma \qquad \ders \Delta} 791 | {\Sigma \ders record\; R \; \Delta : s \leadsto \Sigma, record\; R\; \Delta : s} 792 | \\ 793 | \ru{\Delta \ders \overline{(\hat{\pi} : T)} \to R\; \Delta : s \qquad 794 | U_i = \Delta \to (x : R\; \Delta) \to T_i[x.\pi_j / \hat{\pi_j}] } 795 | {\Sigma[record\; R\; \Delta : s] \ders record\; R\; \Delta\; where\; constructor\; c, \overline{field\; \hat{\pi} : T} \\ \leadsto \Sigma[record\; R\; \Delta : s\; where\; c : \Delta \to \overline{(\pi : T)} \to R\; \Delta; projections\; \overline{\pi : U}]} x \not\in dom(\Delta) 796 | \end{gather*} 797 | 798 | 799 | 800 | 801 | \section{Evaluation} 802 | \label{sec:evaluation} 803 | 804 | 805 | Matching \fbox{$\Sigma \vdash e/q = \sigma$}\ . 806 | \begin{gather*} 807 | \ru{}{\Sigma \vdash v/\lfloor u \rfloor = \cdot} % dot pattern 808 | \qquad 809 | \ru{}{\Sigma \vdash v/x = [v/x]} % var pat 810 | \qquad 811 | \ru 812 | {\Sigma \vdash \overline{v}/\overline{\sigma} = \sigma} 813 | {\Sigma \vdash c\overline{v}/c\overline{p} = \sigma} % con pat 814 | \qquad 815 | \ru{v \bang \pi_i / p_i = \sigma} 816 | {\Sigma \vdash v / \cpi \overline{p} = \sigma} % record pat 817 | \qquad 818 | \ru{}{\Sigma \vdash .\pi/.\pi = \cdot} % proj pat 819 | \\ 820 | \ru{\Sigma \vdash v \to v' \qquad \Sigma \vdash v'/c\overline{p} = \sigma} 821 | {\Sigma \vdash v / c \overline{p} = \sigma} % reduce 822 | \end{gather*} 823 | % \[ 824 | % \begin{array}{ll} 825 | % \overline{e}/\overline{q} = \sigma & 826 | % \Sigma \vdash e/q = \sigma 827 | % \end{array} 828 | % \] 829 | 830 | % \[ 831 | % \begin{array}{ll} 832 | % \inferrule{} 833 | % {\Sigma \vdash v/\lfloor u \rfloor = \cdot} % dot pattern 834 | % & 835 | % \inferrule{} 836 | % {\Sigma \vdash v/x = [v/x]} % var pat 837 | % \\ 838 | % \inferrule{\Sigma \vdash \overline{v}/\overline{\sigma} = \sigma} 839 | % {\Sigma \vdash c\overline{v}/c\overline{p} = \sigma} % con pat 840 | % & 841 | % \inferrule{} 842 | % {\Sigma \vdash .\pi/.\pi = \cdot} % proj pat 843 | % \\ 844 | % \inferrule{\Sigma \vdash v \to v' \qquad \Sigma \vdash v'/c\overline{p} = \sigma} 845 | % {\Sigma \vdash v / c \overline{p} = \sigma} % reduce 846 | % & 847 | % \inferrule{v \bang \pi_i / p_i = \sigma} 848 | % {\Sigma \vdash v / c_{\overline \pi} \overline{p} = \sigma} % record pat 849 | % \end{array} 850 | % \] 851 | 852 | Weak head reduction \fbox{$\Sigma \vdash t \to t'$} 853 | \[ 854 | \begin{array}{lll} 855 | \inferrule{f \overline q = t \in \Sigma \qquad \Sigma \vdash \overline{e}/\overline{q} = \sigma} 856 | {\Sigma \vdash f \overline{e} \overline{e}' \to t[\sigma] \bang \overline{e}'} 857 | & 858 | \ru{} 859 | {\Sigma \vdash \cdot/\cdot = \cdot} %empty 860 | & 861 | \inferrule{\Sigma \vdash e/q = \sigma \qquad \Sigma \vdash \overline{e}/\overline{q} = \sigma'} 862 | {\Sigma \vdash e, \overline{e}/q, \overline{q} = \sigma \uplus \sigma'}%patterns 863 | \end{array} 864 | \] 865 | 866 | 867 | 868 | \section{Coverage} 869 | \label{sec:coverage} 870 | 871 | Pattern refinement 872 | \fbox{$\Gamma \der \vec p : \Delta \prefine x c \Gamma' \der \vec p' : \Delta$}. 873 | \begin{gather*} 874 | \ru{(x : D\,\vec u\,\vec v \in \Gamma) \qquad \Sigma \der D.c : T 875 | }{\Gamma \der \vec p : \Delta \prefine x c \Gamma' \der p \sigma [c\,\vec y / x] : ?} 876 | \end{gather*} 877 | 878 | \section{Termination} 879 | \label{sec:termination} 880 | 881 | \section{Positivity} 882 | \label{sec:positivity} 883 | 884 | \section{Extensions} 885 | \label{sec:extensions} 886 | 887 | \subsection{Extended record declarations} 888 | 889 | Record types 890 | \begin{verbatim} 891 | record R Γ : s where 892 | [constructor c] 893 | [(co)inductive] 894 | [(no-)eta-equality] 895 | rd* 896 | \end{verbatim} 897 | 898 | Record declaration 899 | \begin{verbatim} 900 | rd ::= field π : T 901 | | ts 902 | | cl 903 | \end{verbatim} 904 | 905 | 906 | %% Acknowledgments 907 | \begin{acks} %% acks environment is optional 908 | %% contents suppressed with 'anonymous' 909 | %% Commands \grantsponsor{}{}{} and 910 | %% \grantnum[]{}{} should be used to 911 | %% acknowledge financial support and will be used by metadata 912 | % extraction tools. 913 | % This material is based upon work supported by the 914 | % \grantsponsor{GS100000001}{National Science 915 | % Foundation}{http://dx.doi.org/10.13039/100000001} under Grant 916 | % No.~\grantnum{GS100000001}{nnnnnnn} and Grant 917 | % No.~\grantnum{GS100000001}{mmmmmmm}. Any opinions, findings, and 918 | % conclusions or recommendations expressed in this material are those 919 | % of the author and do not necessarily reflect the views of the 920 | % National Science Foundation. 921 | \end{acks} 922 | 923 | 924 | %% Bibliography 925 | %\bibliography{bibfile} 926 | 927 | 928 | %% Appendix 929 | % \appendix 930 | % \section{Appendix} 931 | 932 | 933 | \end{document} 934 | -------------------------------------------------------------------------------- /src/macros.tex: -------------------------------------------------------------------------------- 1 | % evergreens 2 | \newcommand{\bla}{\ensuremath{\mbox{$$}}} % invisible, but not ignored 3 | \newcommand{\der}{\,\vdash} 4 | \newcommand{\of}{\!:\!} 5 | \newcommand{\is}{\!=\!} 6 | \newcommand{\red}{\longrightarrow} 7 | \newcommand{\restrict}{\upharpoonright} 8 | \newcommand{\FV}{\ensuremath{\mathsf{FV}}} 9 | \newcommand{\NN}{\mathbb{N}} 10 | \newcommand{\defas}{\mathrel{\ :\Longleftrightarrow\ }} 11 | \newcommand{\defiff}{\mathrel{:\Longleftrightarrow}} 12 | \DeclareMathOperator{\dom}{dom} 13 | 14 | % latin etc. abbrev 15 | \newcommand{\abbrev}[1]{#1} % alternative: \emph{#1} 16 | \newcommand{\cf}{\abbrev{cf.}\ } 17 | \newcommand{\eg}{\abbrev{e.\,g.}} 18 | \newcommand{\Eg}{\abbrev{E.\,g.}} 19 | \newcommand{\ie}{\abbrev{i.\,e.}} 20 | \newcommand{\Ie}{\abbrev{I.\,e.}} 21 | \newcommand{\etal}{\abbrev{et.\,al.}} 22 | \newcommand{\wwlog}{w.\,l.\,o.\,g.} % \wlog is ``write into log file'' 23 | \newcommand{\Wlog}{W.\,l.\,o.\,g.} 24 | \newcommand{\wrt}{w.\,r.\,t.} 25 | 26 | % paragraphs 27 | \newcommand{\para}[1]{\paragraph*{\it#1}} 28 | \newcommand{\paradot}[1]{\para{#1.}} 29 | 30 | % proof by cases 31 | \newenvironment{caselist}{% 32 | \begin{list}{{\it Case}}{% 33 | %\setlength{\topsep}{2ex}% DOES NOT SEEM TO WORK 34 | %\setlength{\itemsep}{2ex}% 35 | \setlength{\itemindent}{-2ex}% 36 | }% 37 | }{\end{list}% 38 | } 39 | \newenvironment{subcaselist}{% 40 | \begin{list}{{\it Subcase}}{}% 41 | }{\end{list}% 42 | } 43 | \newenvironment{subsubcaselist}{% 44 | \begin{list}{{\it Subsubcase}}{}% 45 | }{\end{list}% 46 | } 47 | 48 | \newcommand{\nextcase}{\item~} 49 | 50 | % meta-level logic 51 | \newcommand{\mfor}{\ \mbox{for}\ } 52 | \newcommand{\mforsome}{\ \mbox{for some}\ } 53 | \newcommand{\mthen}{\ \mbox{then}\ } 54 | \newcommand{\mif}{\ \mbox{if}\ } 55 | \newcommand{\miff}{\ \mbox{iff}\ } 56 | \newcommand{\motherwise}{\ \mbox{otherwise}} 57 | \newcommand{\mundefined}{\mbox{undefined}} 58 | \newcommand{\mnot}{\mbox{not}\ } 59 | \newcommand{\mand}{\ \mbox{and}\ } 60 | \newcommand{\mor}{\ \mbox{or}\ } 61 | \newcommand{\mimplies}{\ \mbox{implies}\ } 62 | \newcommand{\mimply}{\ \mbox{imply}\ } 63 | \newcommand{\mforall}{\ \mbox{for all}\ } 64 | \newcommand{\mexists}{\mbox{exists}\ } 65 | \newcommand{\mexist}{\mbox{exist}\ } 66 | \newcommand{\mtrue}{\mbox{true}} 67 | \newcommand{\mwith}{\ \mbox{with}\ } 68 | \newcommand{\mholds}{\ \mbox{holds}\ } 69 | \newcommand{\munless}{\ \mbox{unless}\ } 70 | \newcommand{\mboth}{\ \mbox{both}\ } 71 | \newcommand{\msuchthat}{\ \mbox{such that}\ } 72 | % proofs 73 | \newcommand{\msince}{\mbox{since}\ } 74 | \newcommand{\mdef}{\mbox{by def.}} 75 | \newcommand{\mass}{\mbox{assumption}} 76 | \newcommand{\mhyp}{\mbox{by hyp.}} 77 | \newcommand{\mlemma}[1]{\mbox{by Lemma~#1}} 78 | \newcommand{\mih}[1][]{\mbox{by ind.hyp.}#1} 79 | \newcommand{\mgoal}[1][]{\mbox{goal\ifthenelse{\equal{#1}{}}{}{~#1}}} 80 | \newcommand{\mby}[1]{\mbox{by #1}} 81 | \newcommand{\minfrule}{\mbox{by inference rule}} 82 | \newcommand{\mempty}{\mbox{ empty}} 83 | 84 | % Inference rules 85 | \newcommand{\rulename}[1]{\ensuremath{\mbox{\sc#1}}} 86 | \newcommand{\rul}[2]{\dfrac{\begin{array}[b]{@{}l@{}} #1 \end{array}}{#2}} 87 | \newcommand{\ru}[2]{\dfrac{\begin{array}[b]{@{}c@{}} #1 \end{array}}{\begin{array}[l]{@{}c@{}} #2 \end{array}}} 88 | \newcommand{\rux}[3]{\ru{#1}{#2}\ #3} 89 | \newcommand{\nru}[3]{#1\ \ru{#2}{#3}} 90 | \newcommand{\nrux}[4]{#1\ \ru{#2}{#3}\ #4} 91 | \newcommand{\dstack}[2]{\begin{array}[b]{c}#1\\#2\end{array}} 92 | \newcommand{\ndru}[4]{#1\ \ru{\dstack{#2}{#3}}{#4}} 93 | \newcommand{\ndrux}[5]{#1\ \ru{\dstack{#2}{#3}}{#4}\ #5} 94 | \newcommand{\lcol}[1]{\multicolumn{1}{@{}l@{}}{{#1}}} 95 | \newcommand{\rcol}[1]{\multicolumn{1}{@{}r@{}}{{#1}}} 96 | 97 | % Substitution and function update 98 | % read ``\subst F X A'' as ``substitute F for X in A'' 99 | %\newcommand{\subst}[3]{#3[#2 := #1]} 100 | %\newcommand{\subst}[3]{[#1/#2]#3} 101 | \newcommand{\subst}[3]{#3[#1/#2]} 102 | % read ``\update \theta X \G'' as update \theta at point X by \G 103 | \newcommand{\update}[3]{#1[#2 \mapsto #3]} 104 | %\newcommand{\update}[3]{#1,#2 \is #3} 105 | 106 | 107 | % Core agda syntax 108 | \newcommand{\funT}[3]{(#1 : #2) \to #3} 109 | \newcommand{\piT}{\funT} 110 | %\newcommand{\piT}[3]{\Pi #1 \of #2.\, #3} 111 | \newcommand{\lam}[1]{\lambda #1.\,} 112 | \newcommand{\Set}{\mathsf{Set}} 113 | \newcommand{\cpi}{c_{\vec\pi}} % record constructor 114 | \newcommand{\cempty}{\ensuremath{\mathord{\cdot}}} 115 | \newcommand{\dotp}[1]{\lfloor#1\rfloor} % inaccessible pattern 116 | \newcommand{\embp}[1]{\lceil#1\rceil} % embedding of patterns into terms 117 | 118 | \newcommand{\absurd}[1]{\text{absurd}\,#1} 119 | \newcommand{\datasig}{\text{data } D\; \Delta : \Delta' \to s \text{ where } \overline{c : T}} 120 | \newcommand{\recsig}{\text{record } R\;\Delta : s \text{ where 121 | $c : T$; } \text{projections } \overline{\pi : T} } 122 | \newcommand{\funsig}{\text{function } f : T \text{ where } \overline{cl}} 123 | 124 | \DeclareMathOperator{\proj}{proj} 125 | 126 | % Variable names 127 | \newcommand{\vts}{\mathit{ts}} 128 | \newcommand{\vcl}{\mathit{cl}} 129 | \newcommand{\vds}{\mathit{ds}} 130 | \newcommand{\vdd}{\mathit{dd}} 131 | \newcommand{\vrs}{\mathit{rs}} 132 | \newcommand{\vrd}{\mathit{rd}} 133 | \newcommand{\vrhs}{\mathit{rhs}} 134 | 135 | % Core agda judgements 136 | \newcommand{\ders}{\der_\Sigma} 137 | \newcommand{\bang}{\mathrel{!}} 138 | \newcommand{\twobang}{\mathrel{!!}} 139 | \newcommand{\apart}{\mathrel{\#}}%{\mathrel{{<}{>}}} 140 | \newcommand{\prefine}[2]{\stackrel{x \mapsto c}\Longrightarrow} 141 | \newcommand{\noinstance}{\not\preceq} 142 | 143 | %%% Local Variables: 144 | %%% mode: latex 145 | %%% TeX-master: "core-agda" 146 | %%% End: 147 | --------------------------------------------------------------------------------