├── efop_prez
├── prez.out
├── prez.snm
├── prez.toc
├── prez.bbl
├── prez.vrb
├── prez.nav
├── prez.aux
├── prez.tex
├── prez.fdb_latexmk
└── prez.fls
├── department_prez
├── prez.snm
├── prez.bbl
├── 20210129-kovacs-andras-presentation.pdf
├── prez.vrb
├── prez.nav
└── prez.tex
├── .gitignore
├── paper.pdf
├── arxivsub
├── arxivsub.zip
└── lipics-v2021.cls
├── README.md
├── agda
├── README.agda
├── Lib.agda
├── TTGUModel.agda
├── Super.agda
├── TTFLModel.agda
└── IRUniverse.agda
├── references.bib
└── lipics-v2021.cls
/efop_prez/prez.out:
--------------------------------------------------------------------------------
1 |
--------------------------------------------------------------------------------
/efop_prez/prez.snm:
--------------------------------------------------------------------------------
1 |
--------------------------------------------------------------------------------
/department_prez/prez.snm:
--------------------------------------------------------------------------------
1 |
--------------------------------------------------------------------------------
/.gitignore:
--------------------------------------------------------------------------------
1 | *~
2 | *.agdai
3 | .stack-work/
4 | *.o
5 | *.hi
--------------------------------------------------------------------------------
/efop_prez/prez.toc:
--------------------------------------------------------------------------------
1 | \beamer@endinputifotherversion {3.36pt}
2 |
--------------------------------------------------------------------------------
/paper.pdf:
--------------------------------------------------------------------------------
https://raw.githubusercontent.com/AndrasKovacs/universes/HEAD/paper.pdf
--------------------------------------------------------------------------------
/efop_prez/prez.bbl:
--------------------------------------------------------------------------------
1 | \begin{thebibliography}{}
2 |
3 | \end{thebibliography}
4 |
--------------------------------------------------------------------------------
/department_prez/prez.bbl:
--------------------------------------------------------------------------------
1 | \begin{thebibliography}{}
2 |
3 | \end{thebibliography}
4 |
--------------------------------------------------------------------------------
/arxivsub/arxivsub.zip:
--------------------------------------------------------------------------------
https://raw.githubusercontent.com/AndrasKovacs/universes/HEAD/arxivsub/arxivsub.zip
--------------------------------------------------------------------------------
/README.md:
--------------------------------------------------------------------------------
1 | # universes
2 |
3 | Sources for the paper ["Generalized Universe Hierarchies and First-Class Universe Levels"](https://arxiv.org/abs/2103.00223).
4 |
--------------------------------------------------------------------------------
/department_prez/20210129-kovacs-andras-presentation.pdf:
--------------------------------------------------------------------------------
https://raw.githubusercontent.com/AndrasKovacs/universes/HEAD/department_prez/20210129-kovacs-andras-presentation.pdf
--------------------------------------------------------------------------------
/efop_prez/prez.vrb:
--------------------------------------------------------------------------------
1 | \frametitle{Megvalósítás}
2 |
3 | Visszavezetjük az összes feature-t egy már ismert feature-re,
4 | az \textbf{induktív-rekurzív} definíciókra.
5 | \vspace{1em}
6 |
7 | Az indukció-rekurzió egy típuselméleti változata a halmazelméleti \textbf{Mahlo kardinálisoknak}.
8 |
9 | \begin{verbatim}
10 | Univ : Set
11 | Nat : Univ
12 | Π : (A : Univ) → (Interp Univ → Univ) → Univ
13 |
14 | Interp : Univ → Set
15 | Interp Nat = ℕ
16 | Interp (Π A B) = (x : Interp A) → Interp (B x)
17 | \end{verbatim}
18 |
19 |
--------------------------------------------------------------------------------
/department_prez/prez.vrb:
--------------------------------------------------------------------------------
1 | \frametitle{Design choices and variations}
2 |
3 | \textbf{How many universes?} Agda/Coq: countably many.
4 | \vspace{1em}
5 |
6 | \textbf{Are universes totally ordered?} Agda/Coq: yes.
7 | \vspace{1em}
8 |
9 | \textbf{What kind of level polymorphism?}. Coq: bounded polymorphism. Agda: no
10 | bounds allowed. Bounded example:
11 | \begin{verbatim}
12 | myId : (l : Level) → l < 3 → (A : Set l) → A → A
13 | myId l p A x = x
14 | \end{verbatim}
15 |
16 | \textbf{What kind of operations are available on levels?} Agda is more liberal than Coq. Example:
17 | \begin{verbatim}
18 | ℕtoLevel : ℕ → Level
19 | \end{verbatim}
20 |
21 |
22 | \textbf{Are universes cumulative} Agda: no. Coq: yes. Cumulativity: whenever \texttt{A : Set i}, we also have \texttt{A : Set (i + 1)}.
23 |
24 |
--------------------------------------------------------------------------------
/agda/README.agda:
--------------------------------------------------------------------------------
1 |
2 | module README where
3 |
4 | {-
5 | Agda supplement to the paper "Generalized Universe Hierarchies and First-Class Universe Levels".
6 |
7 | This module was checked using Agda 2.6.1, with standard library version 1.3.
8 |
9 | For instructions on installing the standard library, see the following:
10 | https://agda.readthedocs.io/en/v2.6.1.3/tools/package-system.html
11 | -}
12 |
13 | import Lib -- library, for equational reasoning and exporting some definitions
14 | import IRUniverse -- semantic inductive-recursive universes over level structures
15 | import TTGUModel -- model of type theory with generalized levels
16 | import TTFLModel -- model of type theory with first-class levels, where levels are strictly the same as internal Nat
17 | import Super -- some notes on Palmgren's super universes and transfinite hierarchies
18 |
--------------------------------------------------------------------------------
/department_prez/prez.nav:
--------------------------------------------------------------------------------
1 | \beamer@endinputifotherversion {3.36pt}
2 | \headcommand {\slideentry {0}{0}{1}{1/1}{}{0}}
3 | \headcommand {\beamer@framepages {1}{1}}
4 | \headcommand {\slideentry {0}{0}{2}{2/2}{}{0}}
5 | \headcommand {\beamer@framepages {2}{2}}
6 | \headcommand {\slideentry {0}{0}{3}{3/3}{}{0}}
7 | \headcommand {\beamer@framepages {3}{3}}
8 | \headcommand {\slideentry {0}{0}{4}{4/4}{}{0}}
9 | \headcommand {\beamer@framepages {4}{4}}
10 | \headcommand {\slideentry {0}{0}{5}{5/5}{}{0}}
11 | \headcommand {\beamer@framepages {5}{5}}
12 | \headcommand {\slideentry {0}{0}{6}{6/6}{}{0}}
13 | \headcommand {\beamer@framepages {6}{6}}
14 | \headcommand {\slideentry {0}{0}{7}{7/7}{}{0}}
15 | \headcommand {\beamer@framepages {7}{7}}
16 | \headcommand {\slideentry {0}{0}{8}{8/8}{}{0}}
17 | \headcommand {\beamer@framepages {8}{8}}
18 | \headcommand {\slideentry {0}{0}{9}{9/9}{}{0}}
19 | \headcommand {\beamer@framepages {9}{9}}
20 | \headcommand {\slideentry {0}{0}{10}{10/10}{}{0}}
21 | \headcommand {\beamer@framepages {10}{10}}
22 | \headcommand {\beamer@partpages {1}{10}}
23 | \headcommand {\beamer@subsectionpages {1}{10}}
24 | \headcommand {\beamer@sectionpages {1}{10}}
25 | \headcommand {\beamer@documentpages {10}}
26 | \headcommand {\def \inserttotalframenumber {9}}
27 |
--------------------------------------------------------------------------------
/efop_prez/prez.nav:
--------------------------------------------------------------------------------
1 | \beamer@endinputifotherversion {3.36pt}
2 | \headcommand {\slideentry {0}{0}{1}{1/1}{}{0}}
3 | \headcommand {\beamer@framepages {1}{1}}
4 | \headcommand {\slideentry {0}{0}{2}{2/2}{}{0}}
5 | \headcommand {\beamer@framepages {2}{2}}
6 | \headcommand {\slideentry {0}{0}{3}{3/3}{}{0}}
7 | \headcommand {\beamer@framepages {3}{3}}
8 | \headcommand {\slideentry {0}{0}{4}{4/4}{}{0}}
9 | \headcommand {\beamer@framepages {4}{4}}
10 | \headcommand {\slideentry {0}{0}{5}{5/5}{}{0}}
11 | \headcommand {\beamer@framepages {5}{5}}
12 | \headcommand {\slideentry {0}{0}{6}{6/6}{}{0}}
13 | \headcommand {\beamer@framepages {6}{6}}
14 | \headcommand {\slideentry {0}{0}{7}{7/7}{}{0}}
15 | \headcommand {\beamer@framepages {7}{7}}
16 | \headcommand {\slideentry {0}{0}{8}{8/8}{}{0}}
17 | \headcommand {\beamer@framepages {8}{8}}
18 | \headcommand {\slideentry {0}{0}{9}{9/9}{}{0}}
19 | \headcommand {\beamer@framepages {9}{9}}
20 | \headcommand {\slideentry {0}{0}{10}{10/10}{}{0}}
21 | \headcommand {\beamer@framepages {10}{10}}
22 | \headcommand {\slideentry {0}{0}{11}{11/11}{}{0}}
23 | \headcommand {\beamer@framepages {11}{11}}
24 | \headcommand {\slideentry {0}{0}{12}{12/12}{}{0}}
25 | \headcommand {\beamer@framepages {12}{12}}
26 | \headcommand {\beamer@partpages {1}{12}}
27 | \headcommand {\beamer@subsectionpages {1}{12}}
28 | \headcommand {\beamer@sectionpages {1}{12}}
29 | \headcommand {\beamer@documentpages {12}}
30 | \headcommand {\def \inserttotalframenumber {11}}
31 |
--------------------------------------------------------------------------------
/agda/Lib.agda:
--------------------------------------------------------------------------------
1 |
2 | module Lib where
3 |
4 | open import Data.Unit using (⊤; tt) public
5 | open import Data.Empty using (⊥; ⊥-elim) public
6 | open import Data.Product renaming (proj₁ to ₁; proj₂ to ₂) public
7 | open import Data.Sum using (_⊎_; inj₁; inj₂) public
8 | open import Function hiding (id; _∘_; _⇔_) public
9 | open import Induction.WellFounded public
10 | open import Relation.Binary public
11 | open import Relation.Binary.PropositionalEquality
12 | hiding (decSetoid; preorder; setoid; [_]) public
13 |
14 | import Level as L
15 | module F where open import Function public
16 |
17 | coe : ∀ {α}{A B : Set α} → A ≡ B → A → B
18 | coe refl x = x
19 |
20 | _&_ = cong; infixl 9 _&_; {-# DISPLAY cong = _&_ #-}
21 | _◾_ = trans; infixr 4 _◾_; {-# DISPLAY trans = _◾_ #-}
22 | _⁻¹ = sym; infix 6 _⁻¹; {-# DISPLAY sym = _⁻¹ #-}
23 |
24 | coe∘ : ∀ {i}{A B C : Set i}(p : B ≡ C)(q : A ≡ B)(a : A)
25 | → coe p (coe q a) ≡ coe (q ◾ p) a
26 | coe∘ refl refl _ = refl
27 |
28 | UIP : ∀ {i}{A : Set i}{x y : A}{p q : x ≡ y} → p ≡ q
29 | UIP {p = refl}{refl} = refl
30 |
31 | -- function extensionality
32 | postulate
33 | ext : ∀{i j}{A : Set i}{B : A → Set j}{f g : (x : A) → B x}
34 | → ((x : A) → f x ≡ g x) → f ≡ g
35 |
36 | exti : ∀{i j}{A : Set i}{B : A → Set j}{f g : ∀ {x} → B x}
37 | → ((x : A) → f {x} ≡ g {x}) → (λ {x} → f {x}) ≡ g
38 |
39 | unAcc : ∀ {α β A R i} → Acc {α}{β}{A} R i → ∀ {j} → R j i → Acc R j
40 | unAcc (acc f) = f
41 |
42 | Acc-prop : ∀ {α β A R i}(p q : Acc {α}{β}{A} R i) → p ≡ q
43 | Acc-prop (acc f) (acc g) = acc & exti λ j → ext λ p → Acc-prop (f p) (g p)
44 |
45 | happly : ∀ {α β}{A : Set α}{B : Set β}{f g : A → B} → f ≡ g → ∀ a → f a ≡ g a
46 | happly refl a = refl
47 |
48 | _⊗_ :
49 | ∀ {α β}{A : Set α}{B : Set β}
50 | {f g : A → B}(p : f ≡ g){a a' : A}(q : a ≡ a')
51 | → f a ≡ g a'
52 | refl ⊗ refl = refl
53 | infixl 8 _⊗_
54 |
55 | data W {α}{β}(A : Set α) (B : A → Set β) : Set (α L.⊔ β) where
56 | sup : (a : A) → (B a → W A B) → W A B
57 |
--------------------------------------------------------------------------------
/efop_prez/prez.aux:
--------------------------------------------------------------------------------
1 | \relax
2 | \providecommand\hyper@newdestlabel[2]{}
3 | \providecommand\HyperFirstAtBeginDocument{\AtBeginDocument}
4 | \HyperFirstAtBeginDocument{\ifx\hyper@anchor\@undefined
5 | \global\let\oldcontentsline\contentsline
6 | \gdef\contentsline#1#2#3#4{\oldcontentsline{#1}{#2}{#3}}
7 | \global\let\oldnewlabel\newlabel
8 | \gdef\newlabel#1#2{\newlabelxx{#1}#2}
9 | \gdef\newlabelxx#1#2#3#4#5#6{\oldnewlabel{#1}{{#2}{#3}}}
10 | \AtEndDocument{\ifx\hyper@anchor\@undefined
11 | \let\contentsline\oldcontentsline
12 | \let\newlabel\oldnewlabel
13 | \fi}
14 | \fi}
15 | \global\let\hyper@last\relax
16 | \gdef\HyperFirstAtBeginDocument#1{#1}
17 | \providecommand*\HyPL@Entry[1]{}
18 | \bibstyle{alpha}
19 | \HyPL@Entry{0<
>}
20 | \@writefile{toc}{\beamer@endinputifotherversion {3.36pt}}
21 | \@writefile{nav}{\beamer@endinputifotherversion {3.36pt}}
22 | \@writefile{nav}{\headcommand {\slideentry {0}{0}{1}{1/1}{}{0}}}
23 | \@writefile{nav}{\headcommand {\beamer@framepages {1}{1}}}
24 | \HyPL@Entry{1<>}
25 | \@writefile{nav}{\headcommand {\slideentry {0}{0}{2}{2/2}{}{0}}}
26 | \@writefile{nav}{\headcommand {\beamer@framepages {2}{2}}}
27 | \HyPL@Entry{2<>}
28 | \@writefile{nav}{\headcommand {\slideentry {0}{0}{3}{3/3}{}{0}}}
29 | \@writefile{nav}{\headcommand {\beamer@framepages {3}{3}}}
30 | \HyPL@Entry{3<>}
31 | \@writefile{nav}{\headcommand {\slideentry {0}{0}{4}{4/4}{}{0}}}
32 | \@writefile{nav}{\headcommand {\beamer@framepages {4}{4}}}
33 | \HyPL@Entry{4<>}
34 | \@writefile{nav}{\headcommand {\slideentry {0}{0}{5}{5/5}{}{0}}}
35 | \@writefile{nav}{\headcommand {\beamer@framepages {5}{5}}}
36 | \HyPL@Entry{5<>}
37 | \@writefile{nav}{\headcommand {\slideentry {0}{0}{6}{6/6}{}{0}}}
38 | \@writefile{nav}{\headcommand {\beamer@framepages {6}{6}}}
39 | \HyPL@Entry{6<>}
40 | \@writefile{nav}{\headcommand {\slideentry {0}{0}{7}{7/7}{}{0}}}
41 | \@writefile{nav}{\headcommand {\beamer@framepages {7}{7}}}
42 | \HyPL@Entry{7<>}
43 | \@writefile{nav}{\headcommand {\slideentry {0}{0}{8}{8/8}{}{0}}}
44 | \@writefile{nav}{\headcommand {\beamer@framepages {8}{8}}}
45 | \HyPL@Entry{8<>}
46 | \@writefile{nav}{\headcommand {\slideentry {0}{0}{9}{9/9}{}{0}}}
47 | \@writefile{nav}{\headcommand {\beamer@framepages {9}{9}}}
48 | \HyPL@Entry{9<>}
49 | \@writefile{nav}{\headcommand {\slideentry {0}{0}{10}{10/10}{}{0}}}
50 | \@writefile{nav}{\headcommand {\beamer@framepages {10}{10}}}
51 | \HyPL@Entry{10<>}
52 | \@writefile{nav}{\headcommand {\slideentry {0}{0}{11}{11/11}{}{0}}}
53 | \@writefile{nav}{\headcommand {\beamer@framepages {11}{11}}}
54 | \HyPL@Entry{11<>}
55 | \@writefile{nav}{\headcommand {\slideentry {0}{0}{12}{12/12}{}{0}}}
56 | \@writefile{nav}{\headcommand {\beamer@framepages {12}{12}}}
57 | \@writefile{nav}{\headcommand {\beamer@partpages {1}{12}}}
58 | \@writefile{nav}{\headcommand {\beamer@subsectionpages {1}{12}}}
59 | \@writefile{nav}{\headcommand {\beamer@sectionpages {1}{12}}}
60 | \@writefile{nav}{\headcommand {\beamer@documentpages {12}}}
61 | \@writefile{nav}{\headcommand {\def \inserttotalframenumber {11}}}
62 |
--------------------------------------------------------------------------------
/agda/TTGUModel.agda:
--------------------------------------------------------------------------------
1 |
2 | {-
3 | Sketch of the IR model of TTGU, over an arbitrary level structure.
4 | -}
5 |
6 | open import IRUniverse
7 | open import Lib
8 | open import Data.Nat using (ℕ; zero; suc)
9 |
10 | module TTGUModel (L : LvlStruct) where
11 |
12 | module IR = IR-Universe L
13 | open IR hiding (Lift; Lift∘)
14 |
15 | -- base category
16 | Con : Set₁
17 | Con = Set
18 |
19 | Sub : Con → Con → Set
20 | Sub Γ Δ = Γ → Δ
21 |
22 | idₛ : ∀ {Γ} → Sub Γ Γ
23 | idₛ γ = γ
24 |
25 | infixr 4 _∘ₛ_
26 | _∘ₛ_ : ∀ {Γ Δ ξ} → Sub Δ ξ → Sub Γ Δ → Sub Γ ξ
27 | σ ∘ₛ δ = λ γ → σ (δ γ)
28 |
29 | ∙ : Con
30 | ∙ = ⊤
31 |
32 | -- family diagram
33 | Ty : Con → Lvl → Set
34 | Ty Γ i = Γ → U i
35 |
36 | infixl 5 _[_]T
37 | _[_]T : ∀ {Γ Δ i} → Ty Δ i → Sub Γ Δ → Ty Γ i
38 | _[_]T A σ γ = A (σ γ)
39 |
40 | Tm : ∀ Γ {i} → Ty Γ i → Set
41 | Tm Γ A = (γ : Γ) → El (A γ)
42 |
43 | infixl 5 _[_]
44 | _[_] : ∀ {Γ Δ i A} → Tm Δ {i} A → (σ : Sub Γ Δ) → Tm Γ (A [ σ ]T)
45 | _[_] t σ = λ γ → t (σ γ)
46 |
47 | infixl 3 _▶_
48 | _▶_ : ∀(Γ : Con){i} → Ty Γ i → Con
49 | Γ ▶ A = Σ Γ (λ γ → El (A γ))
50 |
51 | p : ∀ {Γ i}(A : Ty Γ i) → Sub (Γ ▶ A) Γ
52 | p A (γ , α) = γ
53 |
54 | q : ∀ {Γ i}(A : Ty Γ i) → Tm (Γ ▶ A) (A [ p A ]T)
55 | q A (γ , α) = α
56 |
57 | infixl 3 _,ₛ_
58 | _,ₛ_ : ∀ {Γ Δ}(σ : Sub Γ Δ){i A} → Tm Γ {i} (_[_]T {i = i} A σ) → Sub Γ (Δ ▶ A)
59 | _,ₛ_ σ t γ = (σ γ) , (t γ)
60 |
61 | Lift : ∀ {Γ i j} → i < j → Ty Γ i → Ty Γ j
62 | Lift p A γ = IR.Lift p (A γ)
63 |
64 | Lift∘ : ∀ {Γ i j k p q A} → Lift {Γ}{j}{k} p (Lift {Γ}{i}{j} q A) ≡ Lift (p ∘ q) A
65 | Lift∘ {p = p} {q} {A} = ext λ γ → IR.Lift∘ q p (A γ)
66 |
67 | lift[]T : ∀ {Γ Δ i j p σ A} → Lift {Δ}{i}{j} p A [ σ ]T ≡ Lift {Γ} p (A [ σ ]T)
68 | lift[]T = refl
69 |
70 | TmLift : ∀ {Γ i j p A} → Tm Γ A ≡ Tm Γ (Lift {Γ}{i}{j} p A)
71 | TmLift {p = p}{A} = (λ f → ∀ x → f x) & ext λ γ → ElLift p (A γ) ⁻¹
72 |
73 | ▶Lift : ∀ {Γ i j A}{p : i < j} → (Γ ▶ A) ≡ (Γ ▶ Lift p A)
74 | ▶Lift {Γ}{A = A}{p} = Σ Γ & ext λ γ → ElLift p (A γ) ⁻¹
75 |
76 | -- term lifting/lowering is definitionally the identity function in ETT
77 | -- also, since semantic term/type formers (with the exception of U) don't depend
78 | -- on levels, lifting/lowering preserves everything.
79 | lift : ∀ {Γ i j}(p : i < j){A : Ty Γ i} → Tm Γ A → Tm Γ (Lift p A)
80 | lift p {A} t = λ γ → coe (ElLift p (A γ) ⁻¹) (t γ)
81 |
82 | lower : ∀ {Γ i j}(p : i < j){A : Ty Γ i} → Tm Γ (Lift p A) → Tm Γ A
83 | lower p {A} t = λ γ → coe (ElLift p (A γ)) (t γ)
84 |
85 | -- universe
86 | Univ : ∀ {Γ i j} → i < j → Ty Γ j
87 | Univ p _ = U' p
88 |
89 | Univ[] : ∀ {Γ Δ σ i j p} → Univ {Δ}{i}{j} p [ σ ]T ≡ Univ {Γ}{i}{j} p
90 | Univ[] = refl
91 |
92 | RussellUniv : ∀ {Γ i j p} → Tm Γ (Univ {Γ}{i}{j} p) ≡ Ty Γ i
93 | RussellUniv = (λ f → ∀ x → f x) & ext λ _ → U<-compute
94 |
95 | -- some type formers
96 | Π : ∀ {Γ i}(A : Ty Γ i) → Ty (Γ ▶ A) i → Ty Γ i
97 | Π {Γ}{i} A B γ = Π' (A γ) λ α → B (γ , α)
98 |
99 | lam : ∀ {Γ i A B} → Tm (Γ ▶ A) B → Tm Γ {i} (Π A B)
100 | lam t γ α = t (γ , α)
101 |
102 | app : ∀ {Γ i A B} → Tm Γ {i} (Π A B) → Tm (Γ ▶ A) B
103 | app t (γ , α) = t γ α
104 |
105 | Nat : ∀ {Γ} i → Ty Γ i
106 | Nat i _ = ℕ' {i}
107 |
108 | LiftNat : ∀ {Γ i j p} → Lift {Γ}{i}{j} p (Nat i) ≡ Nat j
109 | LiftNat = refl
110 |
111 | Zero : ∀ {Γ i} → Tm Γ (Nat i)
112 | Zero _ = zero
113 |
114 | Suc : ∀ {Γ i} → Tm Γ (Nat i) → Tm Γ (Nat i)
115 | Suc t γ = suc (t γ)
116 |
117 | liftZero : ∀ {Γ i j p} → lift {Γ}{i}{j} p {Nat i} (Zero {Γ}{i}) ≡ (Zero {Γ}{j})
118 | liftZero = refl
119 |
120 | liftSuc : ∀ {Γ i j p t} → lift {Γ}{i}{j} p {Nat i} (Suc {Γ}{i} t) ≡ Suc {Γ}{j} (lift {Γ}{i}{j} p {Nat i} t)
121 | liftSuc = refl
122 |
--------------------------------------------------------------------------------
/agda/Super.agda:
--------------------------------------------------------------------------------
1 |
2 | {-# OPTIONS --postfix-projections #-}
3 |
4 | {-
5 | Notes on Palmgren's super universe: http://www2.math.uu.se/~palmgren/universe.pdf
6 |
7 | We investigate recursive sub-universes and transfinite levels.
8 |
9 | Summary of super and Mahlo universes:
10 | (U, El) is *super* if every internal family is contained in an internal sub-universe
11 | (U, El) is *Mahlo* if for every internal (Famᵘ → Famᵘ) operator there is an internal sub-universe closed under it.
12 | -}
13 |
14 | module Super where
15 |
16 | open import Lib
17 | open import Data.Nat
18 | open import Function
19 |
20 | -- simple tree ordinals (Brouwer ordinals) for indexing purposes
21 | data O : Set where
22 | zero : O
23 | suc : O → O
24 | sup : (ℕ → O) → O
25 |
26 | ℕ→O : ℕ → O
27 | ℕ→O zero = zero
28 | ℕ→O (suc n) = suc (ℕ→O n)
29 |
30 | ω : O
31 | ω = sup ℕ→O
32 |
33 | module Cumulative where
34 |
35 | -- universe operator
36 | module _ {U : Set}{El : U → Set} where
37 |
38 | data UU : Set
39 | EL : UU → Set
40 |
41 | data UU where
42 | U' : UU
43 | El' : U → UU
44 | Π' : ∀ a → (El a → UU) → UU
45 | ℕ' : UU
46 |
47 | EL U' = U
48 | EL (El' a) = El a
49 | EL (Π' a b) = ∀ x → EL (b x)
50 | EL ℕ' = ℕ
51 |
52 | -- super universe
53 | data V : Set
54 | S : V → Set
55 |
56 | data V where
57 | UU' : (a : V) → (S a → V) → V
58 | EL' : ∀ {a b} → S (UU' a b) → V
59 | ℕ' : V
60 | Π' : (a : V) → (S a → V) → V
61 | Σ' : (a : V) → (S a → V) → V
62 |
63 | S (UU' a b) = UU {S a} {S ∘ b}
64 | S (EL' x) = EL x
65 | S ℕ' = ℕ
66 | S (Π' a b) = ∀ x → S (b x)
67 | S (Σ' a b) = ∃ (S ∘ b)
68 |
69 | VFam : Set
70 | VFam = ∃ λ a → S a → V
71 |
72 | û : VFam → VFam
73 | û (a , b) = UU' a b , EL' {a}{b}
74 |
75 | û^ : ℕ → VFam
76 | û^ zero = ℕ' , (λ _ → ℕ')
77 | û^ (suc n) = û (û^ n)
78 |
79 | -- codes for some transfinite universes
80 | Uω : V
81 | Uω = UU' (Σ' ℕ' λ n → û^ n .₁) (λ {(n , a) → û^ n .₂ a})
82 |
83 | Uω' : V
84 | Uω' = UU' ℕ' (₁ ∘ û^)
85 |
86 | -- cumulative, non-recursive sub-universes
87 | module NatU where
88 | U : ℕ → Set
89 | El : ∀ {n} → U n → Set
90 | U zero = ⊥
91 | U (suc n) = UU {U n}{El {n}}
92 | El {suc n} a = EL a
93 |
94 | ↑ : ∀ {n} → U n → U (suc n)
95 | ↑ {suc n} a = El' a
96 |
97 | ↑El : ∀ {n a} → El (↑ {n} a) ≡ El a -- not recursive: ↑ ℕ' ≢ ℕ'
98 | ↑El {suc n} {a} = refl
99 |
100 |
101 | -- cumulative, recursive sub-universes (but not transfinite)
102 | module NatU2 where
103 |
104 | U : ℕ → Set
105 | El : ∀ {n} → U n → Set
106 | U zero = ⊥
107 | U (suc n) = UU {U n}{El {n}}
108 | El {suc n} U' = U n
109 | El {suc n} (El' a) = El a
110 | El {suc n} (Π' a b) = ∀ x → El (b x)
111 | El {suc n} ℕ' = ℕ
112 |
113 | ↑ : ∀ {n} → U n → U (suc n)
114 | El≡ : ∀ {n} a → El a ≡ El (↑ {n} a)
115 | ↑ {suc n} U' = El' U'
116 | ↑ {suc n} (El' a) = El' (↑ a)
117 | ↑ {suc n} (Π' a b) = Π' (↑ a) λ x → ↑ (b (coe (El≡ a ⁻¹) x))
118 | ↑ {suc n} ℕ' = ℕ'
119 | El≡ {suc n} U' = refl
120 | El≡ {suc n} (El' a) = El≡ a
121 | El≡ {suc n} (Π' a b) rewrite El≡ a =
122 | (λ b → (∀ x → b x)) & (_∋_ ((λ x → El (b x)) ≡ (λ x → El (↑ (b x)))) (ext λ x → El≡ (b x)))
123 | El≡ {suc n} ℕ' = refl
124 |
125 | -- transfinite, but non-recursive
126 | module OU where
127 |
128 | U : O → Set
129 | El : ∀ {i} → U i → Set
130 | U zero = ⊥
131 | U (suc i) = UU {U i}{El {i}}
132 | U (sup f) = UU {∃ (U ∘ f)}{El ∘ ₂}
133 | El {suc i} a = EL a
134 | El {sup f} a = EL a
135 |
136 | foo : U (suc (suc (suc zero)))
137 | foo = Π' U' λ A → Π' (El' A) λ _ → El' (El' A)
138 |
139 | ex1 : U ω
140 | ex1 = El' (3 , foo)
141 |
142 | ex2 : U ω
143 | ex2 = Π' (3 , foo) λ _ → ℕ'
144 |
145 | ↑+ : ∀ {i} → U i → U (suc i)
146 | ↑+ {i} a = El' a
147 |
148 | ↑sup : ∀ {f n} → U (f n) → U (sup f)
149 | ↑sup {f} {n} a = El' (n , a)
150 |
151 | --------------------------------------------------------------------------------
152 |
--------------------------------------------------------------------------------
/agda/TTFLModel.agda:
--------------------------------------------------------------------------------
1 |
2 | module TTFLModel where
3 |
4 | {-
5 | This file contains a sketch model of the TTFL model when we have ω levels, and
6 | levels are strictly identified with internal natural numbers.
7 | -}
8 |
9 | open import Lib
10 | open import IRUniverse
11 | open import Data.Nat using (ℕ; zero; suc)
12 |
13 | -- First, we construct a concrete level structure with Lvl = ℕ
14 | --------------------------------------------------------------------------------
15 |
16 | ℕ-elim : ∀ {α}(P : ℕ → Set α) → P zero → (∀ {n} → P n → P (suc n)) → ∀ n → P n
17 | ℕ-elim P z s zero = z
18 | ℕ-elim P z s (suc n) = s (ℕ-elim P z s n)
19 |
20 | infix 3 _<*_
21 | _<*_ : ℕ → ℕ → Set
22 | _<*_ a b = ℕ-elim _ ⊥ (λ {b} hyp → (a ≡ b) ⊎ hyp) b
23 |
24 | <*-pred : ∀ i j → suc i <* suc j → i <* j
25 | <*-pred i .(suc i) (inj₁ refl) = inj₁ refl
26 | <*-pred i (suc j) (inj₂ p) = inj₂ (<*-pred i j p)
27 |
28 | <*-irrefl : ∀ i → i <* i → ⊥
29 | <*-irrefl (suc i) p = <*-irrefl i (<*-pred _ _ p)
30 |
31 | lem : ∀ i j → i <* j → i ≢ j
32 | lem i (suc .i) (inj₁ refl) ()
33 | lem .(suc j) (suc j) (inj₂ p) refl = <*-irrefl (suc j) (inj₂ p)
34 |
35 | <-prop* : ∀ i j (p q : i <* j) → p ≡ q
36 | <-prop* i (suc j) (inj₁ p) (inj₁ q) = inj₁ & UIP
37 | <-prop* i (suc j) (inj₁ p) (inj₂ q) = ⊥-elim (lem _ _ q p)
38 | <-prop* i (suc j) (inj₂ p) (inj₁ q) = ⊥-elim (lem _ _ p q)
39 | <-prop* i (suc j) (inj₂ p) (inj₂ q) = inj₂ & <-prop* i j p q
40 |
41 | <*-comp : {i j k : ℕ} → j <* k → i <* j → i <* k
42 | <*-comp {i} {.k} {suc k} (inj₁ refl) q = inj₂ q
43 | <*-comp {i} {j} {suc k} (inj₂ p) q = inj₂ (<*-comp {i}{j}{k} p q)
44 |
45 | postulate
46 | <*-wf : {i : ℕ} → Acc _<*_ i
47 |
48 | L : LvlStruct
49 | L = record {
50 | Lvl = ℕ ;
51 | _<_ = _<*_ ;
52 | <-prop = λ {i}{j}{p}{q} → <-prop* i j p q;
53 | _∘_ = λ {i}{j}{k} → <*-comp {i}{j}{k};
54 | wf = <*-wf }
55 |
56 | -- instantiate IR universe with ℕ
57 | module IR = IR-Universe L
58 | open IR hiding (Lvl'; _<'_; Lift; Lift∘)
59 |
60 | -- Model of TTFL
61 | --------------------------------------------------------------------------------
62 |
63 |
64 | -- base category
65 | Con : Set₁
66 | Con = Set
67 |
68 | Sub : Con → Con → Set
69 | Sub Γ Δ = Γ → Δ
70 |
71 | idₛ : ∀ {Γ} → Sub Γ Γ
72 | idₛ γ = γ
73 |
74 | infixr 4 _∘ₛ_
75 | _∘ₛ_ : ∀ {Γ Δ ξ} → Sub Δ ξ → Sub Γ Δ → Sub Γ ξ
76 | σ ∘ₛ δ = λ γ → σ (δ γ)
77 |
78 | ∙ : Con
79 | ∙ = ⊤
80 |
81 | -- level structure
82 | Level : Con → Set
83 | Level Γ = Γ → Lvl
84 |
85 | Lt : ∀ {Γ} → Level Γ → Level Γ → Set
86 | Lt {Γ} i j = (γ : Γ) → i γ < j γ
87 |
88 | infixl 5 _[_]L
89 | _[_]L : ∀ {Γ Δ} → Level Δ → Sub Γ Δ → Level Γ
90 | _[_]L σ i γ = σ (i γ)
91 |
92 | infixr 4 _∘<_
93 | _∘<_ : ∀ {Γ i j k} → Lt {Γ} j k → Lt i j → Lt i k
94 | _∘<_ {Γ}{i}{j}{k} p q γ = _∘_ {i γ} {j γ}{k γ} (p γ) (q γ)
95 |
96 | infixl 5 _[_]<
97 | _[_]< : ∀ {Γ Δ i j} → Lt i j → (σ : Sub Γ Δ) → Lt (i [ σ ]L) (j [ σ ]L)
98 | _[_]< p σ γ = p (σ γ)
99 |
100 | \thr@@\@toodeep\else
53 | \advance\@enumdepth\@ne
54 | \edef\@enumctr{enum\romannumeral\the\@enumdepth}%
55 | \expandafter
56 | \list
57 | \csname label\@enumctr\endcsname
58 | {\advance\partopsep\topsep
59 | \topsep\z@\@plus\p@
60 | \ifnum\@listdepth=\@ne
61 | \labelsep0.72em
62 | \else
63 | \ifnum\@listdepth=\tw@
64 | \labelsep0.3em
65 | \else
66 | \labelsep0.5em
67 | \fi
68 | \fi
69 | \usecounter\@enumctr\def\makelabel##1{\hss\llap{##1}}}%
70 | \fi}
71 | \def\endenumerate{\ifnum\@listdepth=\@ne\global\@nobotseplisttrue\fi\endlist}
72 | \def\itemize{%
73 | \ifnum \@itemdepth >\thr@@\@toodeep\else
74 | \advance\@itemdepth\@ne
75 | \edef\@itemitem{labelitem\romannumeral\the\@itemdepth}%
76 | \expandafter
77 | \list
78 | \csname\@itemitem\endcsname
79 | {\advance\partopsep\topsep
80 | \topsep\z@\@plus\p@
81 | \ifnum\@listdepth=\@ne
82 | \labelsep0.83em
83 | \else
84 | \ifnum\@listdepth=\tw@
85 | \labelsep0.75em
86 | \else
87 | \labelsep0.5em
88 | \fi
89 | \fi
90 | \def\makelabel##1{\hss\llap{##1}}}%
91 | \fi}
92 | \def\enditemize{\ifnum\@listdepth=\@ne\global\@nobotseplisttrue\fi\endlist}
93 | \def\@title{\textcolor{red}{Author: Please provide a title}}
94 | \def\@sect#1#2#3#4#5#6[#7]#8{%
95 | \ifnum #2>\c@secnumdepth
96 | \let\@svsec\@empty
97 | \else
98 | \refstepcounter{#1}%
99 | \protected@edef\@svsec{\@seccntformat{#1}\relax}%
100 | \fi
101 | \@tempskipa #5\relax
102 | \ifdim \@tempskipa>\z@
103 | \begingroup
104 | #6{%
105 | \@hangfrom{\hskip #3\relax
106 | \ifnum #2=1
107 | \colorbox{lipicsYellow}{\kern0.15em\@svsec\kern0.15em}\quad
108 | \else
109 | \@svsec\quad
110 | \fi}%
111 | \interlinepenalty \@M #8\@@par}%
112 | \endgroup
113 | \csname #1mark\endcsname{#7}%
114 | \addcontentsline{toc}{#1}{%
115 | \ifnum #2>\c@secnumdepth \else
116 | \protect\numberline{\csname the#1\endcsname}%
117 | \fi
118 | #7}%
119 | \else
120 | \def\@svsechd{%
121 | #6{\hskip #3\relax
122 | \@svsec #8}%
123 | \csname #1mark\endcsname{#7}%
124 | \addcontentsline{toc}{#1}{%
125 | \ifnum #2>\c@secnumdepth \else
126 | \protect\numberline{\csname the#1\endcsname}%
127 | \fi
128 | #7}}%
129 | \fi
130 | \@xsect{#5}}
131 | \def\@seccntformat#1{\csname the#1\endcsname}
132 | \def\@biblabel#1{\textcolor{lipicsGray}{\sffamily\bfseries#1}}
133 | \def\EventLogoHeight{25}
134 | \def\copyrightline{%
135 | \ifx\@hideLIPIcs\@undefined
136 | \ifx\@EventLogo\@empty
137 | \else
138 | \setbox\@tempboxa\hbox{\includegraphics[height=\EventLogoHeight\p@]{\@EventLogo}}%
139 | \rlap{\hspace\textwidth\hspace{-\wd\@tempboxa}\hspace{\z@}%
140 | \vtop to\z@{\vskip-0mm\unhbox\@tempboxa\vss}}%
141 | \fi
142 | \scriptsize
143 | \vtop{\hsize\textwidth
144 | \nobreakspace\par
145 | \@Copyright
146 | \ifx\@EventLongTitle\@empty\else\@EventLongTitle.\\\fi
147 | \ifx\@EventEditors\@empty\else
148 | \@Eds: \@EventEditors
149 | ; Article~No.\,\@ArticleNo; pp.\,\@ArticleNo:\thepage--\@ArticleNo:\number\numexpr\getpagerefnumber{TotPages}%
150 | \\
151 | \fi
152 | \setbox\@tempboxa\hbox{\IfFileExists{lipics-logo-bw.pdf}{\includegraphics[height=14\p@,trim=0 15 0 0]{lipics-logo-bw}}{\includegraphics[height=14\p@, width=62pt]{example-image-plain}}}%
153 | \hspace*{\wd\@tempboxa}\enskip
154 | \href{https://www.dagstuhl.de/lipics/}%
155 | {Leibniz International Proceedings in Informatics}\\
156 | \smash{\unhbox\@tempboxa}\enskip
157 | \href{https://www.dagstuhl.de}%
158 | {Schloss Dagstuhl -- Leibniz-Zentrum f{\"u}r Informatik, Dagstuhl Publishing, Germany}}%
159 | \fi}
160 | \def\ps@plain{\let\@mkboth\@gobbletwo
161 | \let\@oddhead\@empty
162 | \let\@evenhead\@empty
163 | \let\@evenfoot\copyrightline
164 | \let\@oddfoot\copyrightline}
165 | \def\lipics@opterrshort{Option "\CurrentOption" not supported}
166 | \def\lipics@opterrlong{The option "\CurrentOption" from article.cls is not supported by lipics.cls.}
167 | \DeclareOption{a5paper}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}}
168 | \DeclareOption{b5paper}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}}
169 | \DeclareOption{legalpaper}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}}
170 | \DeclareOption{executivepaper}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}}
171 | \DeclareOption{landscape}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}}
172 | \DeclareOption{10pt}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}}
173 | \DeclareOption{11pt}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}}
174 | \DeclareOption{12pt}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}}
175 | \DeclareOption{oneside}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}}
176 | \DeclareOption{twoside}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}}
177 | \DeclareOption{titlepage}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}}
178 | \DeclareOption{notitlepage}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}}
179 | \DeclareOption{onecolumn}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}}
180 | \DeclareOption{twocolumn}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}}
181 | \DeclareOption{fleqn}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}}
182 | \DeclareOption{openbib}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}}
183 | \DeclareOption{a4paper}{\PassOptionsToClass{\CurrentOption}{article}
184 | \advance\hoffset-2.95mm
185 | \advance\voffset8.8mm}
186 | \DeclareOption{numberwithinsect}{\let\numberwithinsect\relax}
187 | \DeclareOption{cleveref}{\let\usecleveref\relax}
188 | \DeclareOption{autoref}{\let\useautoref\relax}
189 | \DeclareOption{anonymous}{\let\authoranonymous\relax}
190 | \DeclareOption{thm-restate}{\let\usethmrestate\relax}
191 | \DeclareOption{authorcolumns}{\let\authorcolumns\relax}
192 | \let\compactauthor\relax
193 | \DeclareOption{oldauthorstyle}{\let\compactauthor\@empty}
194 | \DeclareOption{compactauthor}{\let\compactauthor\relax}
195 | \DeclareOption{pdfa}{\let\pdfa\relax}
196 | \DeclareOption*{\PassOptionsToClass{\CurrentOption}{article}}
197 | \ProcessOptions
198 | \LoadClass[twoside,notitlepage,fleqn]{article}
199 | \renewcommand\normalsize{%
200 | \@setfontsize\normalsize\@xpt{13}%
201 | \abovedisplayskip 10\p@ \@plus2\p@ \@minus5\p@
202 | \abovedisplayshortskip \z@ \@plus3\p@
203 | \belowdisplayshortskip 6\p@ \@plus3\p@ \@minus3\p@
204 | \belowdisplayskip \abovedisplayskip
205 | \let\@listi\@listI}
206 | \normalsize
207 | \renewcommand\small{%
208 | \@setfontsize\small\@ixpt{11.5}%
209 | \abovedisplayskip 8.5\p@ \@plus3\p@ \@minus4\p@
210 | \abovedisplayshortskip \z@ \@plus2\p@
211 | \belowdisplayshortskip 4\p@ \@plus2\p@ \@minus2\p@
212 | \def\@listi{\leftmargin\leftmargini
213 | \topsep 4\p@ \@plus2\p@ \@minus2\p@
214 | \parsep 2\p@ \@plus\p@ \@minus\p@
215 | \itemsep \parsep}%
216 | \belowdisplayskip \abovedisplayskip
217 | }
218 | \renewcommand\footnotesize{%
219 | \@setfontsize\footnotesize{8.5}{9.5}%
220 | \abovedisplayskip 6\p@ \@plus2\p@ \@minus4\p@
221 | \abovedisplayshortskip \z@ \@plus\p@
222 | \belowdisplayshortskip 3\p@ \@plus\p@ \@minus2\p@
223 | \def\@listi{\leftmargin\leftmargini
224 | \topsep 3\p@ \@plus\p@ \@minus\p@
225 | \parsep 2\p@ \@plus\p@ \@minus\p@
226 | \itemsep \parsep}%
227 | \belowdisplayskip \abovedisplayskip
228 | }
229 | \renewcommand\large{\@setfontsize\large{10.5}{13}}
230 | \renewcommand\Large{\@setfontsize\Large{12}{14}}
231 | \setlength\parindent{1.5em}
232 | \setlength\headheight{3mm}
233 | \setlength\headsep {10mm}
234 | \setlength\footskip{3mm}
235 | \setlength\textwidth{140mm}
236 | \setlength\textheight{222mm}
237 | \setlength\oddsidemargin{32mm}
238 | \setlength\evensidemargin{38mm}
239 | \setlength\marginparwidth{25mm}
240 | \setlength\topmargin{13mm}
241 | \setlength{\skip\footins}{2\baselineskip \@plus 4\p@ \@minus 2\p@}
242 | \def\@listi{\leftmargin\leftmargini
243 | \parsep\z@ \@plus\p@
244 | \topsep 8\p@ \@plus2\p@ \@minus4\p@
245 | \itemsep \parsep}
246 | \let\@listI\@listi
247 | \@listi
248 | \def\@listii {\leftmargin\leftmarginii
249 | \labelwidth\leftmarginii
250 | \advance\labelwidth-\labelsep
251 | \topsep 4\p@ \@plus2\p@ \@minus\p@
252 | \parsep\z@ \@plus\p@
253 | \itemsep \parsep}
254 | \def\@listiii{\leftmargin\leftmarginiii
255 | \labelwidth\leftmarginiii
256 | \advance\labelwidth-\labelsep
257 | \topsep 2\p@ \@plus\p@\@minus\p@
258 | \parsep \z@
259 | \partopsep \p@ \@plus\z@ \@minus\p@
260 | \itemsep \z@ \@plus\p@}
261 | \def\ps@headings{%
262 | \def\@evenhead{\large\sffamily\bfseries
263 | \llap{\hbox to0.5\oddsidemargin{ \ifx\@hideLIPIcs\@undefined\ifx\@ArticleNo\@empty\textcolor{red}{XX}\else\@ArticleNo\fi:\fi\thepage\hss}}\leftmark\hfil}%
264 | \def\@oddhead{\large\sffamily\bfseries\rightmark\hfil
265 | \rlap{\hbox to0.5\oddsidemargin{\hss \ifx\@hideLIPIcs\@undefined\ifx\@ArticleNo\@empty\textcolor{red}{XX}\else\@ArticleNo\fi:\fi\thepage}}}%
266 | \def\@oddfoot{\hfil
267 | \rlap{%
268 | \vtop{%
269 | \vskip10mm
270 | \colorbox{lipicsYellow}
271 | {\@tempdima\evensidemargin
272 | \advance\@tempdima1in
273 | \advance\@tempdima\hoffset
274 | \hb@xt@\@tempdima{%
275 | \ifx\@hideLIPIcs\@undefined
276 | \textcolor{lipicsGray}{\normalsize\sffamily
277 | \bfseries\quad
278 | \expandafter\textsolittle
279 | \expandafter{\@EventShortTitle}}%
280 | \fi
281 | \strut\hss}}}}}
282 | \let\@evenfoot\@empty
283 | \let\@mkboth\markboth
284 | \let\sectionmark\@gobble
285 | \let\subsectionmark\@gobble}
286 | \pagestyle{headings}
287 | \renewcommand\maketitle{\par
288 | \begingroup
289 | \thispagestyle{plain}
290 | \renewcommand\thefootnote{\@fnsymbol\c@footnote}%
291 | \if@twocolumn
292 | \ifnum \col@number=\@ne
293 | \@maketitle
294 | \else
295 | \twocolumn[\@maketitle]%
296 | \fi
297 | \else
298 | \newpage
299 | \global\@topnum\z@ % Prevents figures from going at top of page.
300 | \@maketitle
301 | \fi
302 | \thispagestyle{plain}\@thanks
303 | \endgroup
304 | \global\let\thanks\relax
305 | \global\let\maketitle\relax
306 | \global\let\@maketitle\relax
307 | \global\let\@thanks\@empty
308 | \global\let\@author\@empty
309 | \global\let\@date\@empty
310 | \global\let\@title\@empty
311 | \global\let\title\relax
312 | \global\let\author\relax
313 | \global\let\date\relax
314 | \global\let\and\relax
315 | }
316 | \newwrite\tocfile
317 | \def\@maketitle{%
318 | \newpage
319 | \null\vskip-\baselineskip
320 | \vskip-\headsep
321 | \@titlerunning
322 | \@authorrunning
323 | %%\let \footnote \thanks
324 | \parindent\z@ \raggedright
325 | \if!\@title!\def\@title{\textcolor{red}{Author: Please fill in a title}}\fi
326 | {\LARGE\sffamily\bfseries\mathversion{bold}\@title \par}%
327 | \vskip 1em
328 | \ifx\@author\orig@author
329 | \textcolor{red}{Author: Please provide author information}%
330 | \else
331 | {\def\thefootnote{\@arabic\c@footnote}%
332 | \setcounter{footnote}{0}%
333 | \fontsize{9.5}{12}\selectfont\@author}%
334 | \fi
335 | \bgroup
336 | \immediate\openout\tocfile=\jobname.vtc
337 | \protected@write\tocfile{
338 | \let\footnote\@gobble
339 | \let\thanks\@gobble
340 | \def\footnotemark{}
341 | \def\and{and }%
342 | \def\,{ }
343 | \def\\{ }
344 | }{%
345 | \string\contitem
346 | \string\title{\@title}%
347 | \string\author{\@authorsfortoc}%
348 | \string\page{\@ArticleNo:\thepage--\@ArticleNo:\number\numexpr\getpagerefnumber{TotPages}}}%
349 | \closeout\tocfile
350 | \egroup
351 | \par}
352 | \renewcommand\tableofcontents{%
353 | \section*{\contentsname}%
354 | \@starttoc{toc}}
355 | \setcounter{secnumdepth}{4}
356 | \renewcommand\section{\@startsection {section}{1}{\z@}%
357 | {-3.5ex \@plus -1ex \@minus -.2ex}%
358 | {2.3ex \@plus.2ex}%
359 | {\sffamily\Large\bfseries\raggedright}}
360 | \renewcommand\subsection{\@startsection{subsection}{2}{\z@}%
361 | {-3.25ex\@plus -1ex \@minus -.2ex}%
362 | {1.5ex \@plus .2ex}%
363 | {\sffamily\Large\bfseries\raggedright}}
364 | \renewcommand\subsubsection{\@startsection{subsubsection}{3}{\z@}%
365 | {-3.25ex\@plus -1ex \@minus -.2ex}%
366 | {1.5ex \@plus .2ex}%
367 | {\sffamily\Large\bfseries\raggedright}}
368 | \renewcommand\paragraph{\@startsection{paragraph}{4}{\z@}%
369 | {-3.25ex \@plus-1ex \@minus-.2ex}%
370 | {1.5ex \@plus .2ex}%
371 | {\sffamily\large\bfseries\raggedright}}
372 | \renewcommand\subparagraph{\@startsection{subparagraph}{5}{\z@}%
373 | {3.25ex \@plus1ex \@minus .2ex}%
374 | {-1em}%
375 | {\sffamily\normalsize\bfseries}}
376 | \newcommand{\proofsubparagraph}{\@startsection{subparagraph}{5}{\z@}%
377 | {3.25ex \@plus1ex \@minus .2ex}%
378 | {-1em}%
379 | {\color{lipicsGray}\sffamily\normalsize\bfseries}}
380 | \setlength\leftmargini \parindent
381 | \setlength\leftmarginii {1.2em}
382 | \setlength\leftmarginiii{1.2em}
383 | \setlength\leftmarginiv {1.2em}
384 | \setlength\leftmarginv {1.2em}
385 | \setlength\leftmarginvi {1.2em}
386 | \renewcommand\labelenumi{%
387 | \textcolor{lipicsGray}{\sffamily\bfseries\upshape\mathversion{bold}\theenumi.}}
388 | \renewcommand\labelenumii{%
389 | \textcolor{lipicsGray}{\sffamily\bfseries\upshape\mathversion{bold}\theenumii.}}
390 | \renewcommand\labelenumiii{%
391 | \textcolor{lipicsGray}{\sffamily\bfseries\upshape\mathversion{bold}\theenumiii.}}
392 | \renewcommand\labelenumiv{%
393 | \textcolor{lipicsGray}{\sffamily\bfseries\upshape\mathversion{bold}\theenumiv.}}
394 | \renewcommand\labelitemi{%
395 | \textcolor{lipicsBulletGray}{\ifnum\@listdepth=\@ne
396 | \rule{0.67em}{0.33em}%
397 | \else
398 | \rule{0.45em}{0.225em}%
399 | \fi}}
400 | \renewcommand\labelitemii{%
401 | \textcolor{lipicsBulletGray}{\rule{0.45em}{0.225em}}}
402 | \renewcommand\labelitemiii{%
403 | \textcolor{lipicsBulletGray}{\sffamily\bfseries\textasteriskcentered}}
404 | \renewcommand\labelitemiv{%
405 | \textcolor{lipicsBulletGray}{\sffamily\bfseries\textperiodcentered}}
406 | \renewenvironment{description}
407 | {\list{}{\advance\partopsep\topsep\topsep\z@\@plus\p@
408 | \labelwidth\z@ \itemindent-\leftmargin
409 | \let\makelabel\descriptionlabel}}
410 | {\ifnum\@listdepth=\@ne\global\@nobotseplisttrue\fi\endlist}
411 | \renewcommand*\descriptionlabel[1]{%
412 | \hspace\labelsep\textcolor{lipicsGray}{\sffamily\bfseries\mathversion{bold}#1}}
413 | \def\topmattervskip{0.7}
414 | \renewenvironment{abstract}{%
415 | \vskip\topmattervskip\bigskipamount
416 | \noindent
417 | \rlap{\color{lipicsLineGray}\vrule\@width\textwidth\@height1\p@}%
418 | \hspace*{7mm}\fboxsep1.5mm\colorbox[rgb]{1,1,1}{\raisebox{-0.4ex}{%
419 | \large\selectfont\sffamily\bfseries\abstractname}}%
420 | \vskip3\p@
421 | \fontsize{9}{12}\selectfont
422 | \noindent\ignorespaces}
423 | {\vskip\topmattervskip\baselineskip\noindent
424 | \subjclassHeading
425 | \ifx\@ccsdescString\@empty
426 | \textcolor{red}{Author: Please fill in 1 or more \string\ccsdesc\space macro}%
427 | \else
428 | \@ccsdescString
429 | \fi
430 | \vskip\topmattervskip\baselineskip
431 | \noindent\keywordsHeading
432 | \ifx\@keywords\@empty
433 | \textcolor{red}{Author: Please fill in \string\keywords\space macro}%
434 | \else
435 | \@keywords
436 | \fi
437 | \ifx\@hideLIPIcs\@undefined
438 | \ifx\@DOIPrefix\@empty\else
439 | \vskip\topmattervskip\baselineskip\noindent
440 | \doiHeading\href{https://doi.org/\@lipicsdoi}{\@lipicsdoi}%
441 | \fi
442 | \fi
443 | \ifx\@category\@empty\else
444 | \vskip\topmattervskip\baselineskip\noindent
445 | \categoryHeading\@category
446 | \fi
447 | \ifx\@relatedversion\@empty\else
448 | \vskip\topmattervskip\baselineskip\noindent
449 | \relatedversionHeading\ifx\authoranonymous\relax\textcolor{red}{Anonymous related version(s)}\else\@relatedversion\fi
450 | \fi
451 | \ifx\@supplement\@empty\else
452 | \vskip\topmattervskip\baselineskip\noindent
453 | \supplementHeading\ifx\authoranonymous\relax\textcolor{red}{Anonymous supplemenatary material}\else\@supplement\fi
454 | \fi
455 | \ifx\@funding\@empty\else
456 | \vskip\topmattervskip\baselineskip\noindent
457 | \fundingHeading\ifx\authoranonymous\relax\textcolor{red}{Anonymous funding}\else\@funding\fi
458 | \fi
459 | \ifx\@acknowledgements\@empty\else
460 | \vskip\topmattervskip\baselineskip\noindent
461 | \acknowledgementsHeading\ifx\authoranonymous\relax\textcolor{red}{Anonymous acknowledgements} \else\@acknowledgements\fi
462 | \fi
463 | \protected@write\@auxout{}{\string\gdef\string\@pageNumberEndAbstract{\thepage}}%
464 | }% end abstract
465 | \renewenvironment{thebibliography}[1]
466 | {\if@noskipsec \leavevmode \fi
467 | \par
468 | \@tempskipa-3.5ex \@plus -1ex \@minus -.2ex\relax
469 | \@afterindenttrue
470 | \@tempskipa -\@tempskipa \@afterindentfalse
471 | \if@nobreak
472 | \everypar{}%
473 | \else
474 | \addpenalty\@secpenalty\addvspace\@tempskipa
475 | \fi
476 | \noindent
477 | \rlap{\color{lipicsLineGray}\vrule\@width\textwidth\@height1\p@}%
478 | \hspace*{7mm}\fboxsep1.5mm\colorbox[rgb]{1,1,1}{\raisebox{-0.4ex}{%
479 | \normalsize\sffamily\bfseries\refname}}%
480 | \@xsect{1ex \@plus.2ex}%
481 | \list{\@biblabel{\@arabic\c@enumiv}}%
482 | {\leftmargin8.5mm
483 | \labelsep\leftmargin
484 | \settowidth\labelwidth{\@biblabel{#1}}%
485 | \advance\labelsep-\labelwidth
486 | \usecounter{enumiv}%
487 | \let\p@enumiv\@empty
488 | \renewcommand\theenumiv{\@arabic\c@enumiv}}%
489 | \fontsize{9}{12}\selectfont
490 | \sloppy
491 | \clubpenalty4000
492 | \@clubpenalty \clubpenalty
493 | \widowpenalty4000%
494 | \sfcode`\.\@m\protected@write\@auxout{}{\string\gdef\string\@pageNumberStartBibliography{\thepage}}}
495 | {\def\@noitemerr
496 | {\@latex@warning{Empty `thebibliography' environment}}%
497 | \protected@write\@auxout{}{\string\gdef\string\@pageNumberEndBibliography{\thepage}}%
498 | \endlist}
499 | \g@addto@macro\appendix{\immediate\write\@auxout{\string\gdef\string\@pageNumberStartAppendix{\thepage}}}%
500 | \renewcommand\footnoterule{%
501 | \kern-8\p@
502 | {\color{lipicsBulletGray}\hrule\@width40mm\@height1\p@}%
503 | \kern6.6\p@}
504 | \renewcommand\@makefntext[1]{%
505 | \parindent\z@\hangindent1em
506 | \leavevmode
507 | \hb@xt@1em{\@makefnmark\hss}#1}
508 | \usepackage{microtype}
509 | \usepackage[utf8]{inputenc}
510 | \IfFileExists{glyphtounicode.tex}{
511 | \input glyphtounicode
512 | \pdfgentounicode=1
513 | }{}%
514 | \IfFileExists{lmodern.sty}{\RequirePackage{lmodern}}{}
515 | \IfFileExists{fontawesome5.sty}{%
516 | \RequirePackage{fontawesome5}%
517 | \IfFileExists{orcid.pdf}{%
518 | \def\orcidsymbol{\includegraphics[height=9\p@]{orcid}}
519 | }{
520 | \def\orcidsymbol{\textcolor{lipicsGray}{\fontsize{9}{12}\sffamily\bfseries \faOrcid}}%
521 | }
522 | \def\mailsymbol{\textcolor{lipicsGray}{\fontsize{9}{12}\sffamily\bfseries \faIcon[regular]{envelope}}}%
523 | \def\homesymbol{\textcolor{lipicsGray}{\fontsize{9}{12}\sffamily\bfseries \faHome}}%
524 | }{%
525 | \ClassWarning{Package fontawesome5 not installed}{Please install package fontawesome5}
526 | \def\orcidsymbol{\textcolor{lipicsGray}{\fontsize{9}{12}\sffamily\bfseries ORCID}}
527 | \def\mailsymbol{\textcolor{lipicsGray}{\fontsize{9}{12}\sffamily\bfseries @}}%
528 | \def\homesymbol{\textcolor{lipicsGray}{\fontsize{9}{12}\sffamily\bfseries H}}%
529 | }%
530 | \RequirePackage[T1]{fontenc}
531 | \RequirePackage{textcomp}
532 | \RequirePackage[mathscr]{eucal}
533 | \RequirePackage{amssymb}
534 | \PassOptionsToPackage{retainmissing}{MnSymbol}
535 | \AtBeginDocument{\@ifpackageloaded{MnSymbol}%
536 | {\expandafter\let\csname ver@amssymb.sty\endcsname\relax
537 | \let\complement\@undefined
538 | \RequirePackage{amssymb}}{}}
539 | \RequirePackage{soul}
540 | \sodef\textsolittle{}{.12em}{.5em\@plus.08em\@minus.06em}%
541 | {.4em\@plus.275em\@minus.183em}
542 | \RequirePackage{color} %kept for backward compatibility
543 | \AtBeginDocument{
544 | \@ifpackageloaded{xcolor}{
545 | }{
546 | \RequirePackage{xcolor}
547 | }
548 | \definecolor{darkgray}{rgb}{0.31,0.31,0.33}
549 | \definecolor[named]{lipicsGray}{rgb}{0.31,0.31,0.33}
550 | \definecolor[named]{lipicsBulletGray}{rgb}{0.60,0.60,0.61}
551 | \definecolor[named]{lipicsLineGray}{rgb}{0.51,0.50,0.52}
552 | \definecolor[named]{lipicsLightGray}{rgb}{0.85,0.85,0.86}
553 | \definecolor[named]{lipicsYellow}{rgb}{0.99,0.78,0.07}
554 | }
555 | \RequirePackage{babel}
556 | \RequirePackage[tbtags,fleqn]{amsmath}
557 | \AtBeginDocument{
558 | \@ifpackageloaded{enumitem}{\ClassWarning{Package 'enuitem' incompatible}{Don't use package 'enumitem'; Package enumerate preloaded!}}{}
559 | \@ifpackageloaded{paralist}{\ClassWarning{Package 'paralist' incompatible}{Don't use package 'paralist'; Package enumerate preloaded!}}{}
560 | }
561 | \RequirePackage{enumerate}
562 | \def\@enum@{\list{\textcolor{lipicsGray}{\sffamily\bfseries\upshape\mathversion{bold}\csname label\@enumctr\endcsname}}%
563 | {\advance\partopsep\topsep
564 | \topsep\z@\@plus\p@
565 | \usecounter{\@enumctr}\def\makelabel##1{\hss\llap{##1}}}}
566 | \def\romanenumerate{\enumerate[(i)]}
567 | \let\endromanenumerate\endenumerate
568 | \def\alphaenumerate{\enumerate[(a)]}
569 | \let\endalphaenumerate\endenumerate
570 | \def\bracketenumerate{\enumerate[(1)]}
571 | \let\endbracketenumerate\endenumerate
572 | \RequirePackage{graphicx}
573 | \RequirePackage{array}
574 | \let\@classzold\@classz
575 | \def\@classz{%
576 | \expandafter\ifx\d@llarbegin\begingroup
577 | \toks \count@ =
578 | \expandafter{\expandafter\small\the\toks\count@}%
579 | \fi
580 | \@classzold}
581 | \RequirePackage{multirow}
582 | \RequirePackage{tabularx}
583 | \RequirePackage[online]{threeparttable}
584 | \def\TPTtagStyle#1{#1)}
585 | \def\tablenotes{\small\TPT@defaults
586 | \@ifnextchar[\TPT@setuptnotes\TPTdoTablenotes} % ]
587 | \RequirePackage{listings}
588 | \lstset{basicstyle=\small\ttfamily,%
589 | backgroundcolor=\color{lipicsLightGray},%
590 | frame=single,framerule=0pt,xleftmargin=\fboxsep,xrightmargin=\fboxsep}
591 | \RequirePackage[left,mathlines]{lineno}
592 | \linenumbers
593 | \renewcommand\linenumberfont{\normalfont\tiny\sffamily}
594 | %%%% patch to cope with amsmath
595 | %%%% http://phaseportrait.blogspot.de/2007/08/lineno-and-amsmath-compatibility.html
596 | \newcommand*\patchAmsMathEnvironmentForLineno[1]{%
597 | \expandafter\let\csname old#1\expandafter\endcsname\csname #1\endcsname
598 | \expandafter\let\csname oldend#1\expandafter\endcsname\csname end#1\endcsname
599 | \renewenvironment{#1}%
600 | {\linenomath\csname old#1\endcsname}%
601 | {\csname oldend#1\endcsname\endlinenomath}}%
602 | \newcommand*\patchBothAmsMathEnvironmentsForLineno[1]{%
603 | \patchAmsMathEnvironmentForLineno{#1}%
604 | \patchAmsMathEnvironmentForLineno{#1*}}%
605 | \AtBeginDocument{%
606 | \patchBothAmsMathEnvironmentsForLineno{equation}%
607 | \patchBothAmsMathEnvironmentsForLineno{align}%
608 | \patchBothAmsMathEnvironmentsForLineno{flalign}%
609 | \patchBothAmsMathEnvironmentsForLineno{alignat}%
610 | \patchBothAmsMathEnvironmentsForLineno{gather}%
611 | \patchBothAmsMathEnvironmentsForLineno{multline}}
612 | \let\usehyperxmp\@empty%
613 | \ifx\pdfa\relax%
614 | \IfFileExists{hyperxmp.sty}{%
615 | \RequirePackage{hyperxmp}%
616 | \@ifpackagelater{hyperxmp}{2019/04/05}{%
617 | \let\usehyperxmp\relax%
618 | }{%
619 | \ClassWarning{Package hyperxmp outdated}{You are using an outdated version of the package hyperxmp. Please update!}%
620 | }}{}%
621 | \fi%
622 | \IfFileExists{totpages.sty}{
623 | \RequirePackage{totpages}
624 | }{
625 | \ClassWarning{Package totpages not installed}{Please install package totpages}
626 | \newcounter{TotPages}
627 | \setcounter{TotPages}{99}
628 | }
629 | \ifx\usehyperxmp\relax%
630 | \RequirePackage[pdfa,unicode]{hyperref}%
631 | \else%
632 | \RequirePackage[unicode]{hyperref}%
633 | \fi%
634 | \let\C\relax%
635 | \let\G\relax%
636 | \let\F\relax%
637 | \let\U\relax%
638 | \pdfstringdefDisableCommands{%
639 | \let\thanks\@gobble%
640 | \let\footnote\@gobble%
641 | \def\footnotemark{}%
642 | \def\cs#1{\textbackslash #1}%
643 | \let\normalfont\@empty%
644 | \let\scshape\@empty%
645 | \def\and{and }%
646 | \def\,{ }%
647 | \def\textrightarrow{ -> }%
648 | \let\mathsf\@empty%
649 | }%
650 | \hypersetup{
651 | breaklinks=true,
652 | pdfencoding=unicode,
653 | bookmarksnumbered,
654 | pdfborder={0 0 0},
655 | pdfauthor={ }
656 | }%
657 | \AtBeginDocument{
658 | \ifx\usehyperxmp\relax
659 | \hypersetup{
660 | pdftitle={\@title},
661 | pdfauthor={\ifx\authoranonymous\relax Anonymous author(s) \else \@authorsforpdf \fi},
662 | pdfkeywords={\@keywords},
663 | pdfproducer={LaTeX with lipics-v2021.cls},
664 | pdfsubject={LIPIcs, Vol.\@SeriesVolume, \@EventShortTitle},
665 | pdfcopyright = { Copyright (C) \ifx\authoranonymous\relax Anonymous author(s) \else \@copyrightholder; \fi licensed under Creative Commons License CC-BY 4.0},
666 | pdflang={en},
667 | pdfmetalang={en},
668 | pdfpublisher={Schloss Dagstuhl -- Leibniz-Zentrum für Informatik, Dagstuhl Publishing, Germany},
669 | pdflicenseurl={https://creativecommons.org/licenses/by/4.0/},
670 | pdfpubtype={LIPIcs},
671 | pdfvolumenum={\@SeriesVolume},
672 | pdfpagerange={\@ArticleNo:\thepage-\@ArticleNo:\theTotPages},
673 | pdfdoi={\@lipicsdoi},
674 | pdfapart=3,
675 | pdfaconformance=B
676 | }
677 | \else%
678 | \hypersetup{
679 | pdftitle={\@title},
680 | pdfauthor={\ifx\authoranonymous\relax Anonymous author(s) \else \@authorsforpdf \fi},
681 | pdfkeywords={\@keywords},
682 | pdfcreator={LaTeX with lipics-v2021.cls},
683 | pdfsubject={LIPIcs, Vol.\@SeriesVolume, \@EventShortTitle; Copyright (C) \ifx\authoranonymous\relax Anonymous author(s) \else \@copyrightholder; \fi licensed under Creative Commons License CC-BY 4.0}
684 | }%
685 | \fi %
686 | }
687 | \ifx\usehyperxmp\relax
688 | \pdfobjcompresslevel=0
689 | \pdfinclusioncopyfonts=1
690 | \IfFileExists{colorprofiles.tex}{
691 | \RequirePackage{colorprofiles}%
692 | \IfFileExists{sRGB.icc}{
693 | \immediate\pdfobj stream attr{/N 3} file{sRGB.icc}
694 | \pdfcatalog{%
695 | /OutputIntents [
696 | <<
697 | /Type /OutputIntent
698 | /S /GTS_PDFA1
699 | /DestOutputProfile \the\pdflastobj\space 0 R
700 | /OutputConditionIdentifier (sRGB)
701 | /Info (sRGB)
702 | >>
703 | ]
704 | }}{}
705 | }{\ClassWarning{Package colorprofiles not installed}{Please install package colorprofiles}}
706 | \fi
707 | \RequirePackage[labelsep=space,singlelinecheck=false,%
708 | font={up,small},labelfont={sf,bf},%
709 | listof=false]{caption}%"listof" instead of "list" for backward compatibility
710 | \@ifpackagelater{hyperref}{2009/12/09}
711 | {\captionsetup{compatibility=false}}%cf. http://groups.google.de/group/comp.text.tex/browse_thread/thread/db9310eb540fbbd8/42e30f3b7b3aa17a?lnk=raot
712 | {}
713 | \DeclareCaptionLabelFormat{boxed}{%
714 | \kern0.05em{\color[rgb]{0.99,0.78,0.07}\rule{0.73em}{0.73em}}%
715 | \hspace*{0.67em}\bothIfFirst{#1}{~}#2}
716 | \captionsetup{labelformat=boxed}
717 | \captionsetup[table]{position=top}
718 | \RequirePackage[figuresright]{rotating}
719 | \caption@AtBeginDocument{\@ifpackageloaded{subfig}{\ClassError{lipics}{%
720 | Do not load the subfig package}{The more recent subcaption package is already loaded}}{}}
721 | \RequirePackage{subcaption}
722 | \def\titlerunning#1{\gdef\@titlerunning{{\let\footnote\@gobble\markboth{#1}{#1}}}}
723 | \def\authorrunning#1{%
724 | \gdef\@authorrunning{\markright{\ifx\authoranonymous\relax\textcolor{red}{Anonymous author(s)} \else\if!#1!\textcolor{red}{Author: Please fill in the \string\authorrunning\space macro}\else#1\fi\fi}}}
725 | \titlerunning{\@title}
726 | \authorrunning{\textcolor{red}{Author: Please use the \string\authorrunning\space macro}}
727 | \def\EventLongTitle#1{\gdef\@EventLongTitle{#1}}
728 | \EventLongTitle{}
729 | \def\EventShortTitle#1{\gdef\@EventShortTitle{#1}}
730 | \EventShortTitle{}
731 | \def\EventEditors#1{\gdef\@EventEditors{#1}}
732 | \EventEditors{}
733 | \def\EventNoEds#1{\gdef\@EventNoEds{#1}\xdef\@Eds{Editor\ifnum#1>1s\fi}}
734 | \EventNoEds{1}
735 | \def\EventLogo#1{\gdef\@EventLogo{#1}}
736 | \EventLogo{}
737 | \def\EventAcronym#1{\gdef\@EventAcronym{#1}}
738 | \EventAcronym{}
739 | \def\EventYear#1{\gdef\@EventYear{#1}}
740 | \EventYear{}
741 | \def\EventDate#1{\gdef\@EventDate{#1}}
742 | \EventDate{}
743 | \def\EventLocation#1{\gdef\@EventLocation{#1}}
744 | \EventLocation{}
745 | \def\SeriesVolume#1{\gdef\@SeriesVolume{#1}}
746 | \SeriesVolume{}
747 | \def\ArticleNo#1{\gdef\@ArticleNo{#1}}
748 | \ArticleNo{}
749 | \def\DOIPrefix#1{\gdef\@DOIPrefix{#1}}
750 | \DOIPrefix{10.4230/LIPIcs}
751 | \def\@lipicsdoi{\@DOIPrefix.\@EventAcronym.\@EventYear.\@ArticleNo}
752 | \def\and{\newline}
753 | \let\orig@author\@author
754 | \let\@authorsfortoc\@empty
755 | \let\@authorsforpdf\@empty
756 | \newcount\c@author
757 | \newcounter{currentauthor}
758 | \def\authorcolumnsMin{6}
759 | \def\@authornum{0}
760 | \def\author#1#2#3#4#5{%
761 | \ifx\@author\orig@author\let\@author\@empty\fi
762 | \g@addto@macro\@author{%
763 | \noexpandarg\StrBehind{#2}{\and \url}[\homepageTemp]\IfSubStr{#2}{\and \url}{\StrBefore{#2}{\and \url}[\affiliation]}{\def\affiliation{#2}}%
764 | \expandarg\exploregroups\StrRemoveBraces{\homepageTemp}[\homepage]%
765 | \ifx\authorcolumns\relax
766 | \ifnum\c@author>\authorcolumnsMin
767 | \stepcounter{currentauthor}
768 | \ifodd\value{currentauthor}
769 | \begin{minipage}[t]{\textwidth}
770 | \begin{minipage}[t]{0.49\textwidth}
771 | \else
772 | \hfill \begin{minipage}[t]{0.49\textwidth}
773 | \fi
774 | \else
775 | \ClassWarning{Option 'authorcolumns' only applicable for > 6 authors}{Option 'authorcolumns' only applicable for >6 authors!}
776 | \addvspace{0.5\baselineskip}
777 | \fi
778 | \else
779 | \addvspace{0.5\baselineskip}
780 | \fi
781 | {\Large\bfseries
782 | \if!#1!
783 | \textcolor{red}{Author: Please enter author name}%
784 | \else
785 | \ifx\authoranonymous\relax
786 | \textcolor{red}{Anonymous author}
787 | \else
788 | #1\,%
789 | \ifx\compactauthor\relax\if!#3!\else{\,\href{mailto:#3}{\mailsymbol}}\fi%
790 | \ifx\homepage\@empty\else{\,\href{\homepage}{\homesymbol}}\fi\fi%
791 | \if!#4!\else{\,\href{#4}{\orcidsymbol}}\fi%
792 | \if!#5!\else
793 | \ifx\@funding\@empty
794 | \expandafter\g@addto@macro\expandafter\@funding{\textit{\expandafter{\let\footnote\@gobble #1}}:\space{#5}}
795 | \else
796 | \expandafter\g@addto@macro\expandafter\@funding{\\\textit{\expandafter{\let\footnote\@gobble #1}}:\space{#5}}
797 | \fi
798 | \fi
799 | \fi
800 | \fi
801 | }
802 | {\small
803 | \if!#2!\textcolor{red}{Author: Please enter affiliation as second parameter of the author macro}\else{\\* \ifx\authoranonymous\relax\textcolor{red}{Anonymous affiliation}\else\ifx\compactauthor\relax \affiliation \else#2\fi\fi}\fi
804 | \ifx\compactauthor\relax\else\if!#3!\else{\ifx\authoranonymous\relax\else\\*\href{mailto:#3}{#3}\fi}\fi\fi
805 | }\par
806 | \ifx\authorcolumns\relax
807 | \ifnum\c@author>\authorcolumnsMin
808 | \end{minipage}
809 | \ifnum\c@author=\value{currentauthor}
810 | \end{minipage}
811 | \else
812 | \ifodd\value{currentauthor}
813 | \else
814 | \end{minipage}%
815 | \medskip
816 | \fi
817 | \fi
818 | \fi
819 | \fi}%
820 | \global\advance\c@author\@ne
821 | \protected@write\@auxout{}{\string\gdef\string\@authornum{\the\c@author}}
822 | \ifnum\c@author=\@ne
823 | \gdef\@authorsfortoc{#1}%
824 | \gdef\@authorsforpdf{#1}
825 | \else
826 | \expandafter\g@addto@macro\expandafter\@authorsforpdf\expandafter{, #1}
827 | \expandafter\g@addto@macro\expandafter\@authorsfortoc\expandafter{\expandafter\csname\the\c@author authand\endcsname#1}%
828 | \@namedef{\the\c@author authand}{,\space}%
829 | \AtBeginDocument{%
830 | \expandafter\ifnum\@authornum=2
831 | \@namedef{2authand}{\space and\space}%
832 | \else
833 | \@namedef{\@authornum authand}{,\space and\space}%
834 | \fi}
835 | \fi}
836 | \newcommand*\affil[2][]{%
837 | \ClassError{lipics}
838 | {\string\affil\space deprecated: Please enter affiliation as second parameter of the author macro}
839 | {Since 2017, \string\affil\space is obsolete in lipics.}}
840 | \newcommand*\Copyright[1]{%
841 | \def\@copyrightholder{#1}
842 | \def\@Copyright{%
843 | \setbox\@tempboxa\hbox{\IfFileExists{cc-by.pdf}{\includegraphics[height=14\p@,clip]{cc-by}}{\includegraphics[height=14\p@, width=40pt]{example-image-plain}}}%
844 | \@rightskip\@flushglue \rightskip\@rightskip
845 | \hangindent\dimexpr\wd\@tempboxa+0.5em\relax
846 | \href{https://creativecommons.org/licenses/by/4.0/}%
847 | {\smash{\lower\baselineskip\hbox{\unhcopy\@tempboxa}}}\enskip
848 | \textcopyright\ %
849 | \ifx!#1!\textcolor{red}{Author: Please fill in the \string\Copyright\space macro}\else\ifx\authoranonymous\relax\textcolor{red}{Anonymous author(s)}\else#1\fi\fi
850 | ;\\%
851 | licensed under Creative Commons License CC-BY 4.0\ifx!#1!\\\null\fi\par}}
852 | \Copyright{\textcolor{red}{Author: Please provide a copyright holder}}
853 | \let\@copyrightholder\@empty
854 | \def\hideLIPIcs{\let\@hideLIPIcs\relax}
855 | \usepackage{xstring}
856 | \def\keywords#1{\def\@keywords{#1}}
857 | \let\@keywords\@empty
858 | \def\keywordsHeading{%
859 | \textcolor{lipicsGray}{\fontsize{9}{12}\sffamily\bfseries
860 | Keywords and phrases\enskip}}
861 | \RequirePackage{comment}
862 | \excludecomment{CCSXML}
863 | % inspired by https://tex.stackexchange.com/questions/12810/how-do-i-split-a-string
864 | \global\newcommand\ccsdesc[2][100]{\@ccsdesc#1~#2~~\relax}
865 | \let\orig@ccsdesc\@ccsdesc
866 | \let\@ccsdesc\@empty
867 | \let\@ccsdescString\@empty
868 | \gdef\@ccsdesc#1~#2~#3~{
869 | \ifx\@ccsdesc\orig@ccsdesc\let\@ccsdesc\@empty\fi
870 | \ifx!#3!
871 | \ifx\@ccsdescString\@empty
872 | \g@addto@macro\@ccsdescString{{#2}}
873 | \else
874 | \g@addto@macro\@ccsdescString{; {#2}}
875 | \fi
876 | \else
877 | \ifx\@ccsdescString\@empty
878 | \g@addto@macro\@ccsdescString{{#2} $\rightarrow$ {#3}}
879 | \else
880 | \g@addto@macro\@ccsdescString{; {#2} $\rightarrow$ {#3}}
881 | \fi
882 | \fi
883 | \ccsdescEnd
884 | }
885 | \def\ccsdescEnd#1\relax{}
886 | \def\subjclass#1{
887 | \ClassError{lipics}
888 | {\string\subjclass\space deprecated: Please enter subject classification in 1 or more ccsdesc macros}
889 | {Since 2019, \string\subjclass\space is obsolete in lipics.}}
890 | \let\@subjclass\@empty
891 | \def\subjclassHeading{%
892 | \textcolor{lipicsGray}{\fontsize{9}{12}\sffamily\bfseries
893 | 2012 ACM Subject Classification\enskip}}
894 | \def\doiHeading{%
895 | \textcolor{lipicsGray}{\fontsize{9}{12}\sffamily\bfseries
896 | Digital Object Identifier\enskip}}
897 | \def\category#1{\def\@category{#1}}
898 | \let\@category\@empty
899 | \def\categoryHeading{%
900 | \textcolor{lipicsGray}{\fontsize{9}{12}\sffamily\bfseries
901 | Category\enskip}}
902 | \def\relatedversion#1{\def\@relatedversion{#1}}
903 | \let\@relatedversion\@empty
904 | \define@key{relatedversiondetails}{linktext}{\def\relatedversiondetails@linktext{#1}}
905 | \define@key{relatedversiondetails}{cite}{\def\relatedversiondetails@cite{#1}}
906 | \newcommand*\addtorelatedversionmacro[2]{%
907 | \ifx\@relatedversion\@empty%
908 | \g@addto@macro\@relatedversion{#1}%
909 | \else%
910 | \g@addto@macro\@relatedversion{\\#1}%
911 | \fi%
912 | }%
913 | \newcommand{\relatedversiondetails}[3][]{%
914 | \begingroup%
915 | \let\relatedversiondetails@linktext\@empty
916 | \let\relatedversiondetails@cite\@empty
917 | \setkeys{relatedversiondetails}{#1}%
918 | \ifx\relatedversiondetails@linktext\@empty%
919 | \protected@edef\tmp{\textit{#2}:\space{\url{#3}}}%
920 | \else%
921 | \protected@edef\tmp{\textit{#2}:\space{\href{#3}{\texttt{\relatedversiondetails@linktext}}}}%
922 | \fi%
923 | \ifx\relatedversiondetails@cite\@empty%
924 | \else%
925 | \protected@edef\tmp{\tmp\space\cite{\relatedversiondetails@cite}}%
926 | \fi%
927 | \expandafter\addtorelatedversionmacro\expandafter{\tmp}{#1}%
928 | \endgroup%
929 | }%
930 | \def\relatedversionHeading{%
931 | \textcolor{lipicsGray}{\fontsize{9}{12}\sffamily\bfseries
932 | Related Version\enskip}}
933 | \def\supplement#1{\def\@supplement{#1}}
934 | \let\@supplement\@empty
935 | \define@key{supplementdetails}{linktext}{\def\supplementdetails@linktext{#1}}
936 | \define@key{supplementdetails}{cite}{\def\supplementdetails@cite{#1}}
937 | \define@key{supplementdetails}{subcategory}{\def\supplementdetails@subcategory{#1}}
938 | \define@key{supplementdetails}{swhid}{\def\supplementdetails@swhid{#1}}
939 | \newcommand*\addtosupplementmacro[2]{%
940 | \ifx\@supplement\@empty%
941 | \g@addto@macro\@supplement{#1}%
942 | \else%
943 | \g@addto@macro\@supplement{\\#1}%
944 | \fi%
945 | }%
946 | \newcommand{\supplementdetails}[3][]{%
947 | \begingroup%
948 | \let\supplementdetails@linktext\@empty
949 | \let\supplementdetails@cite\@empty
950 | \let\supplementdetails@subcategory\@empty
951 | \let\supplementdetails@swhid\@empty
952 | \setkeys{supplementdetails}{#1}%
953 | \ifx\supplementdetails@subcategory\@empty%
954 | \protected@edef\tmp{\textit{#2}}
955 | \else
956 | \protected@edef\tmp{\textit{#2\,\,(\supplementdetails@subcategory)}}%
957 | \fi
958 | \ifx\supplementdetails@linktext\@empty%
959 | \protected@edef\tmp{\tmp:\space{\url{#3}}}%
960 | \else%
961 | \protected@edef\tmp{\tmp:\space{\href{#3}{\texttt{\supplementdetails@linktext}}}}%
962 | \fi%
963 | \ifx\supplementdetails@cite\@empty%
964 | \else%
965 | \protected@edef\tmp{\tmp\space\cite{\supplementdetails@cite}}%
966 | \fi
967 | \ifx\supplementdetails@swhid\@empty%
968 | \else%
969 | \protected@edef\tmp{\tmp \\ \hspace*{1.2em} archived at \href{https://archive.softwareheritage.org/\supplementdetails@swhid}{\texttt{\supplementdetails@swhid}}}%
970 | \fi%
971 | \expandafter\addtosupplementmacro\expandafter{\tmp}{#1}%
972 | \endgroup%
973 | }%
974 | \def\supplementHeading{%
975 | \textcolor{lipicsGray}{\fontsize{9}{12}\sffamily\bfseries
976 | Supplementary Material\enskip}}
977 | \newcommand\flag[2][0.9cm]{%
978 | \leavevmode\marginpar{%
979 | \raisebox{\dimexpr-\totalheight+\ht\strutbox\relax}%
980 | [\dimexpr\ht\strutbox+3mm][\dp\strutbox]{\expandafter\includegraphics[width=#1]{#2}}%
981 | }}
982 | \def\funding#1{\def\@funding{#1}}
983 | \let\@funding\@empty
984 | \def\fundingHeading{%
985 | \textcolor{lipicsGray}{\fontsize{9}{12}\sffamily\bfseries
986 | Funding\enskip}}
987 | \def\acknowledgements#1{\def\@acknowledgements{#1}}
988 | \let\@acknowledgements\@empty
989 | \def\acknowledgementsHeading{%
990 | \textcolor{lipicsGray}{\fontsize{9}{12}\sffamily\bfseries
991 | Acknowledgements\enskip}}
992 | \RequirePackage{amsthm}
993 | \ifx\usethmrestate\relax
994 | \RequirePackage{thm-restate}
995 | \fi
996 | \thm@headfont{%
997 | \textcolor{lipicsGray}{$\blacktriangleright$}\nobreakspace\sffamily\bfseries}
998 | \def\th@remark{%
999 | \thm@headfont{%
1000 | \textcolor{lipicsGray}{$\blacktriangleright$}\nobreakspace\sffamily}%
1001 | \normalfont % body font
1002 | \thm@preskip\topsep \divide\thm@preskip\tw@
1003 | \thm@postskip\thm@preskip
1004 | }
1005 | \def\@endtheorem{\endtrivlist}%\@endpefalse
1006 | \renewcommand\qedsymbol{\textcolor{lipicsGray}{\ensuremath{\blacktriangleleft}}}
1007 | \renewenvironment{proof}[1][\proofname]{\par
1008 | \pushQED{\qed}%
1009 | \normalfont \topsep6\p@\@plus6\p@\relax
1010 | \trivlist
1011 | \item[\hskip\labelsep
1012 | \color{lipicsGray}\sffamily\bfseries
1013 | #1\@addpunct{.}]\ignorespaces
1014 | }{%
1015 | \popQED\endtrivlist%\@endpefalse
1016 | }
1017 | \newcommand{\claimqedhere}{\renewcommand\qedsymbol{\textcolor{lipicsGray}{\ensuremath{\vartriangleleft}}}%
1018 | \qedhere%
1019 | \renewcommand\qedsymbol{\textcolor{lipicsGray}{\ensuremath{\blacktriangleleft}}}}
1020 | \newenvironment{claimproof}[1][\proofname]{
1021 | \pushQED{\qed}%
1022 | \normalfont \topsep6\p@\@plus6\p@\relax
1023 | \trivlist
1024 | \item[\hskip\labelsep
1025 | \color{lipicsGray}\sffamily
1026 | #1\@addpunct{.}]\ignorespaces
1027 | }{%
1028 | \renewcommand\qedsymbol{\textcolor{lipicsGray}{\ensuremath{\vartriangleleft}}}
1029 | \popQED\endtrivlist%\@endpefalse
1030 | \renewcommand\qedsymbol{\textcolor{lipicsGray}{\ensuremath{\blacktriangleleft}}}
1031 | }
1032 | % inspired by qed of amsthm class
1033 | \DeclareRobustCommand{\lipicsEnd}{%
1034 | \leavevmode\unskip\penalty9999 \hbox{}\nobreak\hfill
1035 | \quad\hbox{$\lrcorner$}%
1036 | }
1037 | \AtBeginDocument{
1038 | \@ifpackageloaded{algorithm2e}{
1039 | \@ifpackagelater{algorithm2e}{2009/11/17}{
1040 | \renewcommand{\algorithmcfname}{\sffamily\bfseries{}Algorithm}%
1041 | \renewcommand{\@algocf@procname}{\sffamily\bfseries{}Procedure}%
1042 | \SetAlgoCaptionSeparator{~}
1043 | \SetAlCapHSkip{0pt}
1044 | \renewcommand{\algocf@captiontext}[2]{%
1045 | \kern0.05em{\color{lipicsYellow}\rule{0.73em}{0.73em}}%
1046 | \hspace*{0.67em}\small #1\algocf@capseparator\nobreakspace#2}
1047 | \renewcommand{\algocf@makecaption}[2]{%
1048 | \parbox[t]{\textwidth}{\algocf@captiontext{#1}{#2}}%
1049 | }%
1050 | \renewcommand{\algocf@captionproctext}[2]{%
1051 | {%
1052 | \kern0.05em{\color{lipicsYellow}\rule{0.73em}{0.73em}}%
1053 | \hspace*{0.67em}\small%
1054 | \ProcSty{\ProcFnt\algocf@procname\ifthenelse{\boolean{algocf@procnumbered}}{\nobreakspace\thealgocf\algocf@typo\algocf@capseparator}{\relax}}%
1055 | \nobreakspace\ProcNameSty{\ProcNameFnt\algocf@captname #2@}% Name of the procedure in ProcName Style.
1056 | \ifthenelse{\equal{\algocf@captparam #2@}{\arg@e}}{}{% if no argument, write nothing
1057 | \ProcNameSty{\ProcNameFnt(}\ProcArgSty{\ProcArgFnt\algocf@captparam #2@}\ProcNameSty{\ProcNameFnt)}%else put arguments in ProcArgSty:
1058 | }% endif
1059 | \algocf@captother #2@%
1060 | }%
1061 | }%
1062 | \renewcommand{\@algocf@capt@boxed}{above}
1063 | \renewcommand{\@algocf@capt@ruled}{above}
1064 | \setlength\algotitleheightrule{0pt}
1065 | }{\ClassWarning{%
1066 | Package algorithm2e outdated}{You are using an outdated version of the package algorithm2e. Please update!}}
1067 | }{}
1068 | \@ifpackageloaded{algorithm}{
1069 | \captionsetup[algorithm]{name=Algorithm, labelformat=boxed, position=top}
1070 | \newcommand\fs@ruled@notop{\def\@fs@cfont{\bfseries}\let\@fs@capt\floatc@ruled
1071 | \def\@fs@pre{}%
1072 | \def\@fs@post{\kern2pt\hrule\relax}%
1073 | \def\@fs@mid{\kern2pt\hrule\kern2pt}%
1074 | \let\@fs@iftopcapt\iftrue}
1075 | \@ifundefined{fst@algorithm}{}{
1076 | \renewcommand\fst@algorithm{\fs@ruled@notop}
1077 | }
1078 | }{}
1079 | \ifx\usecleveref\relax\else
1080 | \@ifpackageloaded{cleveref}{\ClassWarning{Use document option 'cleveref' instead}{Use document option 'cleveref' instead directly loading package 'cleveref'}}{}
1081 | \fi
1082 | \ifx\usethmrestate\relax\else
1083 | \@ifpackageloaded{thm-restate}{\ClassWarning{Use document option 'thm-restate' instead}{Use document option 'thm-restate' instead directly loading package 'thm-restate'}}{}
1084 | \fi
1085 | \ifx\useautoref\relax
1086 | \@ifundefined{algorithmautorefname}{\newcommand{\algorithmautorefname}{Algorithm}}{\renewcommand{\algorithmautorefname}{Algorithm}}%
1087 | \fi
1088 | }
1089 |
1090 | \ifx\usecleveref\relax
1091 | \RequirePackage[capitalise, noabbrev]{cleveref}
1092 | \crefname{algocf}{Algorithm}{Algorithms}
1093 | \Crefname{algocf}{Algorithm}{Algorithms}
1094 | \newcommand{\crefrangeconjunction}{--}
1095 | \newcommand{\creflastconjunction}{, and\nobreakspace}
1096 | \fi
1097 | \ifx\useautoref\relax
1098 | \RequirePackage{aliascnt}
1099 | \fi
1100 | \newtheoremstyle{claimstyle}{\topsep}{\topsep}{}{0pt}{\sffamily}{. }{5pt plus 1pt minus 1pt}%
1101 | {$\vartriangleright$ \thmname{#1}\thmnumber{ #2}\thmnote{ (#3)}}
1102 | \theoremstyle{plain}
1103 | \newtheorem{theorem}{Theorem}
1104 | \ifx\numberwithinsect\relax
1105 | \@addtoreset{theorem}{section}
1106 | \expandafter\def\expandafter\thetheorem\expandafter{%
1107 | \expandafter\thesection\expandafter\@thmcountersep\thetheorem}
1108 | \fi
1109 |
1110 | \ifx\useautoref\relax
1111 | \addto\extrasenglish{%
1112 | \def\chapterautorefname{Chapter}%
1113 | \def\sectionautorefname{Section}%
1114 | \def\subsectionautorefname{Subsection}%
1115 | \def\subsubsectionautorefname{Subsubsection}%
1116 | \def\paragraphautorefname{Paragraph}%
1117 | \def\subparagraphautorefname{Subparagraph}%
1118 | }
1119 | \addto\extrasUKenglish{%
1120 | \def\chapterautorefname{Chapter}%
1121 | \def\sectionautorefname{Section}%
1122 | \def\subsectionautorefname{Subsection}%
1123 | \def\subsubsectionautorefname{Subsubsection}%
1124 | \def\paragraphautorefname{Paragraph}%
1125 | \def\subparagraphautorefname{Subparagraph}%
1126 | }
1127 | \addto\extrasUSenglish{%
1128 | \def\chapterautorefname{Chapter}%
1129 | \def\sectionautorefname{Section}%
1130 | \def\subsectionautorefname{Subsection}%
1131 | \def\subsubsectionautorefname{Subsubsection}%
1132 | \def\paragraphautorefname{Paragraph}%
1133 | \def\subparagraphautorefname{Subparagraph}%
1134 | }
1135 | \ifx\usethmrestate\relax
1136 | \newtheorem{lemma}[theorem]{Lemma}
1137 | \newtheorem{corollary}[theorem]{Corollary}
1138 | \newtheorem{proposition}[theorem]{Proposition}
1139 | \newtheorem{exercise}[theorem]{Exercise}
1140 | \newtheorem{definition}[theorem]{Definition}
1141 | \newtheorem{conjecture}[theorem]{Conjecture}
1142 | \newtheorem{observation}[theorem]{Observation}
1143 | \theoremstyle{definition}
1144 | \newtheorem{example}[theorem]{Example}
1145 | \theoremstyle{remark}
1146 | \newtheorem{note}[theorem]{Note}
1147 | \newtheorem*{note*}{Note}
1148 | \newtheorem{remark}[theorem]{Remark}
1149 | \newtheorem*{remark*}{Remark}
1150 | \theoremstyle{claimstyle}
1151 | \newtheorem{claim}[theorem]{Claim}
1152 | \newtheorem*{claim*}{Claim}
1153 | \else
1154 | \newaliascnt{lemma}{theorem}
1155 | \newtheorem{lemma}[lemma]{Lemma}
1156 | \aliascntresetthe{lemma}
1157 | \newcommand{\lemmaautorefname}{Lemma}
1158 | \newaliascnt{corollary}{theorem}
1159 | \newtheorem{corollary}[corollary]{Corollary}
1160 | \aliascntresetthe{corollary}
1161 | \newcommand{\corollaryautorefname}{Corollary}
1162 | \newaliascnt{proposition}{theorem}
1163 | \newtheorem{proposition}[proposition]{Proposition}
1164 | \aliascntresetthe{proposition}
1165 | \newcommand{\propositionautorefname}{Proposition}
1166 | \newaliascnt{exercise}{theorem}
1167 | \newtheorem{exercise}[exercise]{Exercise}
1168 | \aliascntresetthe{exercise}
1169 | \newcommand{\exerciseautorefname}{Exercise}
1170 | \newaliascnt{definition}{theorem}
1171 | \newtheorem{definition}[definition]{Definition}
1172 | \aliascntresetthe{definition}
1173 | \newcommand{\definitionautorefname}{Definition}
1174 | \newaliascnt{conjecture}{theorem}
1175 | \newtheorem{conjecture}[conjecture]{Conjecture}
1176 | \aliascntresetthe{conjecture}
1177 | \newcommand{\conjectureautorefname}{Conjecture}
1178 | \newaliascnt{observation}{theorem}
1179 | \newtheorem{observation}[observation]{Observation}
1180 | \aliascntresetthe{observation}
1181 | \newcommand{\observationautorefname}{Observation}
1182 | \theoremstyle{definition}
1183 | \newaliascnt{example}{theorem}
1184 | \newtheorem{example}[example]{Example}
1185 | \aliascntresetthe{example}
1186 | \newcommand{\exampleautorefname}{Example}
1187 | \theoremstyle{remark}
1188 | \newaliascnt{note}{theorem}
1189 | \newtheorem{note}[note]{Note}
1190 | \aliascntresetthe{note}
1191 | \newcommand{\noteautorefname}{Note}
1192 | \newtheorem*{note*}{Note}
1193 | \newaliascnt{remark}{theorem}
1194 | \newtheorem{remark}[remark]{Remark}
1195 | \aliascntresetthe{remark}
1196 | \newcommand{\remarkautorefname}{Remark}
1197 | \newtheorem*{remark*}{Remark}
1198 | \theoremstyle{claimstyle}
1199 | \newaliascnt{claim}{theorem}
1200 | \newtheorem{claim}[claim]{Claim}
1201 | \aliascntresetthe{claim}
1202 | \newcommand{\claimautorefname}{Claim}
1203 | \newtheorem*{claim*}{Claim}
1204 | \fi
1205 | \else
1206 | \newtheorem{lemma}[theorem]{Lemma}
1207 | \newtheorem{corollary}[theorem]{Corollary}
1208 | \newtheorem{proposition}[theorem]{Proposition}
1209 | \newtheorem{exercise}[theorem]{Exercise}
1210 | \newtheorem{definition}[theorem]{Definition}
1211 | \newtheorem{conjecture}[theorem]{Conjecture}
1212 | \newtheorem{observation}[theorem]{Observation}
1213 | \theoremstyle{definition}
1214 | \newtheorem{example}[theorem]{Example}
1215 | \theoremstyle{remark}
1216 | \newtheorem{note}[theorem]{Note}
1217 | \newtheorem*{note*}{Note}
1218 | \newtheorem{remark}[theorem]{Remark}
1219 | \newtheorem*{remark*}{Remark}
1220 | \theoremstyle{claimstyle}
1221 | \newtheorem{claim}[theorem]{Claim}
1222 | \newtheorem*{claim*}{Claim}
1223 | \fi
1224 | \theoremstyle{plain}
1225 | \endinput
1226 | %%
1227 | %% End of file `lipics-v2021.cls'.
1228 |
--------------------------------------------------------------------------------
/arxivsub/lipics-v2021.cls:
--------------------------------------------------------------------------------
1 | %%
2 | %% This is file `lipics-v2021.cls'.
3 | %%
4 | %% -----------------------------------------------------------------
5 | %% Author: Dagstuhl Publishing & le-tex publishing services
6 | %%
7 | %% This file is part of the lipics package for preparing
8 | %% LIPICS articles.
9 | %%
10 | %% Copyright (C) 2020 Schloss Dagstuhl
11 | %%
12 | %% This work may be distributed and/or modified under the
13 | %% conditions of the LaTeX Project Public License, either version 1.3
14 | %% of this license or (at your option) any later version.
15 | %% The latest version of this license is in
16 | %% http://www.latex-project.org/lppl.txt
17 | %% and version 1.3 or later is part of all distributions of LaTeX
18 | %% version 2005/12/01 or later.
19 | %%
20 | %% This work has the LPPL maintenance status `maintained'.
21 | %%
22 | %% The Current Maintainer of this work is
23 | %% Schloss Dagstuhl (publishing@dagstuhl.de).
24 | %% -----------------------------------------------------------------
25 | %%
26 | \ProvidesClass{lipics-v2021}
27 | [2021/01/04 v3.1.0 LIPIcs articles]
28 | \NeedsTeXFormat{LaTeX2e}[2015/01/01]
29 | \emergencystretch1em
30 | \advance\hoffset-1in
31 | \advance\voffset-1in
32 | \advance\hoffset2.95mm
33 | \newif\if@nobotseplist \@nobotseplistfalse
34 | \def\@endparenv{%
35 | \addpenalty\@endparpenalty\if@nobotseplist\else\addvspace\@topsepadd\fi\@endpetrue}
36 | \def\@doendpe{%
37 | \@endpetrue
38 | \def\par{\@restorepar
39 | \everypar{}%
40 | \par
41 | \if@nobotseplist
42 | \addvspace\topsep
43 | \addvspace\partopsep
44 | \global\@nobotseplistfalse
45 | \fi
46 | \@endpefalse}%
47 | \everypar{{\setbox\z@\lastbox}%
48 | \everypar{}%
49 | \if@nobotseplist\global\@nobotseplistfalse\fi
50 | \@endpefalse}}
51 | \def\enumerate{%
52 | \ifnum \@enumdepth >\thr@@\@toodeep\else
53 | \advance\@enumdepth\@ne
54 | \edef\@enumctr{enum\romannumeral\the\@enumdepth}%
55 | \expandafter
56 | \list
57 | \csname label\@enumctr\endcsname
58 | {\advance\partopsep\topsep
59 | \topsep\z@\@plus\p@
60 | \ifnum\@listdepth=\@ne
61 | \labelsep0.72em
62 | \else
63 | \ifnum\@listdepth=\tw@
64 | \labelsep0.3em
65 | \else
66 | \labelsep0.5em
67 | \fi
68 | \fi
69 | \usecounter\@enumctr\def\makelabel##1{\hss\llap{##1}}}%
70 | \fi}
71 | \def\endenumerate{\ifnum\@listdepth=\@ne\global\@nobotseplisttrue\fi\endlist}
72 | \def\itemize{%
73 | \ifnum \@itemdepth >\thr@@\@toodeep\else
74 | \advance\@itemdepth\@ne
75 | \edef\@itemitem{labelitem\romannumeral\the\@itemdepth}%
76 | \expandafter
77 | \list
78 | \csname\@itemitem\endcsname
79 | {\advance\partopsep\topsep
80 | \topsep\z@\@plus\p@
81 | \ifnum\@listdepth=\@ne
82 | \labelsep0.83em
83 | \else
84 | \ifnum\@listdepth=\tw@
85 | \labelsep0.75em
86 | \else
87 | \labelsep0.5em
88 | \fi
89 | \fi
90 | \def\makelabel##1{\hss\llap{##1}}}%
91 | \fi}
92 | \def\enditemize{\ifnum\@listdepth=\@ne\global\@nobotseplisttrue\fi\endlist}
93 | \def\@title{\textcolor{red}{Author: Please provide a title}}
94 | \def\@sect#1#2#3#4#5#6[#7]#8{%
95 | \ifnum #2>\c@secnumdepth
96 | \let\@svsec\@empty
97 | \else
98 | \refstepcounter{#1}%
99 | \protected@edef\@svsec{\@seccntformat{#1}\relax}%
100 | \fi
101 | \@tempskipa #5\relax
102 | \ifdim \@tempskipa>\z@
103 | \begingroup
104 | #6{%
105 | \@hangfrom{\hskip #3\relax
106 | \ifnum #2=1
107 | \colorbox{lipicsYellow}{\kern0.15em\@svsec\kern0.15em}\quad
108 | \else
109 | \@svsec\quad
110 | \fi}%
111 | \interlinepenalty \@M #8\@@par}%
112 | \endgroup
113 | \csname #1mark\endcsname{#7}%
114 | \addcontentsline{toc}{#1}{%
115 | \ifnum #2>\c@secnumdepth \else
116 | \protect\numberline{\csname the#1\endcsname}%
117 | \fi
118 | #7}%
119 | \else
120 | \def\@svsechd{%
121 | #6{\hskip #3\relax
122 | \@svsec #8}%
123 | \csname #1mark\endcsname{#7}%
124 | \addcontentsline{toc}{#1}{%
125 | \ifnum #2>\c@secnumdepth \else
126 | \protect\numberline{\csname the#1\endcsname}%
127 | \fi
128 | #7}}%
129 | \fi
130 | \@xsect{#5}}
131 | \def\@seccntformat#1{\csname the#1\endcsname}
132 | \def\@biblabel#1{\textcolor{lipicsGray}{\sffamily\bfseries#1}}
133 | \def\EventLogoHeight{25}
134 | \def\copyrightline{%
135 | \ifx\@hideLIPIcs\@undefined
136 | \ifx\@EventLogo\@empty
137 | \else
138 | \setbox\@tempboxa\hbox{\includegraphics[height=\EventLogoHeight\p@]{\@EventLogo}}%
139 | \rlap{\hspace\textwidth\hspace{-\wd\@tempboxa}\hspace{\z@}%
140 | \vtop to\z@{\vskip-0mm\unhbox\@tempboxa\vss}}%
141 | \fi
142 | \scriptsize
143 | \vtop{\hsize\textwidth
144 | \nobreakspace\par
145 | \@Copyright
146 | \ifx\@EventLongTitle\@empty\else\@EventLongTitle.\\\fi
147 | \ifx\@EventEditors\@empty\else
148 | \@Eds: \@EventEditors
149 | ; Article~No.\,\@ArticleNo; pp.\,\@ArticleNo:\thepage--\@ArticleNo:\number\numexpr\getpagerefnumber{TotPages}%
150 | \\
151 | \fi
152 | \setbox\@tempboxa\hbox{\IfFileExists{lipics-logo-bw.pdf}{\includegraphics[height=14\p@,trim=0 15 0 0]{lipics-logo-bw}}{\includegraphics[height=14\p@, width=62pt]{example-image-plain}}}%
153 | \hspace*{\wd\@tempboxa}\enskip
154 | \href{https://www.dagstuhl.de/lipics/}%
155 | {Leibniz International Proceedings in Informatics}\\
156 | \smash{\unhbox\@tempboxa}\enskip
157 | \href{https://www.dagstuhl.de}%
158 | {Schloss Dagstuhl -- Leibniz-Zentrum f{\"u}r Informatik, Dagstuhl Publishing, Germany}}%
159 | \fi}
160 | \def\ps@plain{\let\@mkboth\@gobbletwo
161 | \let\@oddhead\@empty
162 | \let\@evenhead\@empty
163 | \let\@evenfoot\copyrightline
164 | \let\@oddfoot\copyrightline}
165 | \def\lipics@opterrshort{Option "\CurrentOption" not supported}
166 | \def\lipics@opterrlong{The option "\CurrentOption" from article.cls is not supported by lipics.cls.}
167 | \DeclareOption{a5paper}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}}
168 | \DeclareOption{b5paper}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}}
169 | \DeclareOption{legalpaper}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}}
170 | \DeclareOption{executivepaper}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}}
171 | \DeclareOption{landscape}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}}
172 | \DeclareOption{10pt}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}}
173 | \DeclareOption{11pt}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}}
174 | \DeclareOption{12pt}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}}
175 | \DeclareOption{oneside}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}}
176 | \DeclareOption{twoside}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}}
177 | \DeclareOption{titlepage}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}}
178 | \DeclareOption{notitlepage}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}}
179 | \DeclareOption{onecolumn}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}}
180 | \DeclareOption{twocolumn}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}}
181 | \DeclareOption{fleqn}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}}
182 | \DeclareOption{openbib}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}}
183 | \DeclareOption{a4paper}{\PassOptionsToClass{\CurrentOption}{article}
184 | \advance\hoffset-2.95mm
185 | \advance\voffset8.8mm}
186 | \DeclareOption{numberwithinsect}{\let\numberwithinsect\relax}
187 | \DeclareOption{cleveref}{\let\usecleveref\relax}
188 | \DeclareOption{autoref}{\let\useautoref\relax}
189 | \DeclareOption{anonymous}{\let\authoranonymous\relax}
190 | \DeclareOption{thm-restate}{\let\usethmrestate\relax}
191 | \DeclareOption{authorcolumns}{\let\authorcolumns\relax}
192 | \let\compactauthor\relax
193 | \DeclareOption{oldauthorstyle}{\let\compactauthor\@empty}
194 | \DeclareOption{compactauthor}{\let\compactauthor\relax}
195 | \DeclareOption{pdfa}{\let\pdfa\relax}
196 | \DeclareOption*{\PassOptionsToClass{\CurrentOption}{article}}
197 | \ProcessOptions
198 | \LoadClass[twoside,notitlepage,fleqn]{article}
199 | \renewcommand\normalsize{%
200 | \@setfontsize\normalsize\@xpt{13}%
201 | \abovedisplayskip 10\p@ \@plus2\p@ \@minus5\p@
202 | \abovedisplayshortskip \z@ \@plus3\p@
203 | \belowdisplayshortskip 6\p@ \@plus3\p@ \@minus3\p@
204 | \belowdisplayskip \abovedisplayskip
205 | \let\@listi\@listI}
206 | \normalsize
207 | \renewcommand\small{%
208 | \@setfontsize\small\@ixpt{11.5}%
209 | \abovedisplayskip 8.5\p@ \@plus3\p@ \@minus4\p@
210 | \abovedisplayshortskip \z@ \@plus2\p@
211 | \belowdisplayshortskip 4\p@ \@plus2\p@ \@minus2\p@
212 | \def\@listi{\leftmargin\leftmargini
213 | \topsep 4\p@ \@plus2\p@ \@minus2\p@
214 | \parsep 2\p@ \@plus\p@ \@minus\p@
215 | \itemsep \parsep}%
216 | \belowdisplayskip \abovedisplayskip
217 | }
218 | \renewcommand\footnotesize{%
219 | \@setfontsize\footnotesize{8.5}{9.5}%
220 | \abovedisplayskip 6\p@ \@plus2\p@ \@minus4\p@
221 | \abovedisplayshortskip \z@ \@plus\p@
222 | \belowdisplayshortskip 3\p@ \@plus\p@ \@minus2\p@
223 | \def\@listi{\leftmargin\leftmargini
224 | \topsep 3\p@ \@plus\p@ \@minus\p@
225 | \parsep 2\p@ \@plus\p@ \@minus\p@
226 | \itemsep \parsep}%
227 | \belowdisplayskip \abovedisplayskip
228 | }
229 | \renewcommand\large{\@setfontsize\large{10.5}{13}}
230 | \renewcommand\Large{\@setfontsize\Large{12}{14}}
231 | \setlength\parindent{1.5em}
232 | \setlength\headheight{3mm}
233 | \setlength\headsep {10mm}
234 | \setlength\footskip{3mm}
235 | \setlength\textwidth{140mm}
236 | \setlength\textheight{222mm}
237 | \setlength\oddsidemargin{32mm}
238 | \setlength\evensidemargin{38mm}
239 | \setlength\marginparwidth{25mm}
240 | \setlength\topmargin{13mm}
241 | \setlength{\skip\footins}{2\baselineskip \@plus 4\p@ \@minus 2\p@}
242 | \def\@listi{\leftmargin\leftmargini
243 | \parsep\z@ \@plus\p@
244 | \topsep 8\p@ \@plus2\p@ \@minus4\p@
245 | \itemsep \parsep}
246 | \let\@listI\@listi
247 | \@listi
248 | \def\@listii {\leftmargin\leftmarginii
249 | \labelwidth\leftmarginii
250 | \advance\labelwidth-\labelsep
251 | \topsep 4\p@ \@plus2\p@ \@minus\p@
252 | \parsep\z@ \@plus\p@
253 | \itemsep \parsep}
254 | \def\@listiii{\leftmargin\leftmarginiii
255 | \labelwidth\leftmarginiii
256 | \advance\labelwidth-\labelsep
257 | \topsep 2\p@ \@plus\p@\@minus\p@
258 | \parsep \z@
259 | \partopsep \p@ \@plus\z@ \@minus\p@
260 | \itemsep \z@ \@plus\p@}
261 | \def\ps@headings{%
262 | \def\@evenhead{\large\sffamily\bfseries
263 | \llap{\hbox to0.5\oddsidemargin{ \ifx\@hideLIPIcs\@undefined\ifx\@ArticleNo\@empty\textcolor{red}{XX}\else\@ArticleNo\fi:\fi\thepage\hss}}\leftmark\hfil}%
264 | \def\@oddhead{\large\sffamily\bfseries\rightmark\hfil
265 | \rlap{\hbox to0.5\oddsidemargin{\hss \ifx\@hideLIPIcs\@undefined\ifx\@ArticleNo\@empty\textcolor{red}{XX}\else\@ArticleNo\fi:\fi\thepage}}}%
266 | \def\@oddfoot{\hfil
267 | \rlap{%
268 | \vtop{%
269 | \vskip10mm
270 | \colorbox{lipicsYellow}
271 | {\@tempdima\evensidemargin
272 | \advance\@tempdima1in
273 | \advance\@tempdima\hoffset
274 | \hb@xt@\@tempdima{%
275 | \ifx\@hideLIPIcs\@undefined
276 | \textcolor{lipicsGray}{\normalsize\sffamily
277 | \bfseries\quad
278 | \expandafter\textsolittle
279 | \expandafter{\@EventShortTitle}}%
280 | \fi
281 | \strut\hss}}}}}
282 | \let\@evenfoot\@empty
283 | \let\@mkboth\markboth
284 | \let\sectionmark\@gobble
285 | \let\subsectionmark\@gobble}
286 | \pagestyle{headings}
287 | \renewcommand\maketitle{\par
288 | \begingroup
289 | \thispagestyle{plain}
290 | \renewcommand\thefootnote{\@fnsymbol\c@footnote}%
291 | \if@twocolumn
292 | \ifnum \col@number=\@ne
293 | \@maketitle
294 | \else
295 | \twocolumn[\@maketitle]%
296 | \fi
297 | \else
298 | \newpage
299 | \global\@topnum\z@ % Prevents figures from going at top of page.
300 | \@maketitle
301 | \fi
302 | \thispagestyle{plain}\@thanks
303 | \endgroup
304 | \global\let\thanks\relax
305 | \global\let\maketitle\relax
306 | \global\let\@maketitle\relax
307 | \global\let\@thanks\@empty
308 | \global\let\@author\@empty
309 | \global\let\@date\@empty
310 | \global\let\@title\@empty
311 | \global\let\title\relax
312 | \global\let\author\relax
313 | \global\let\date\relax
314 | \global\let\and\relax
315 | }
316 | \newwrite\tocfile
317 | \def\@maketitle{%
318 | \newpage
319 | \null\vskip-\baselineskip
320 | \vskip-\headsep
321 | \@titlerunning
322 | \@authorrunning
323 | %%\let \footnote \thanks
324 | \parindent\z@ \raggedright
325 | \if!\@title!\def\@title{\textcolor{red}{Author: Please fill in a title}}\fi
326 | {\LARGE\sffamily\bfseries\mathversion{bold}\@title \par}%
327 | \vskip 1em
328 | \ifx\@author\orig@author
329 | \textcolor{red}{Author: Please provide author information}%
330 | \else
331 | {\def\thefootnote{\@arabic\c@footnote}%
332 | \setcounter{footnote}{0}%
333 | \fontsize{9.5}{12}\selectfont\@author}%
334 | \fi
335 | \bgroup
336 | \immediate\openout\tocfile=\jobname.vtc
337 | \protected@write\tocfile{
338 | \let\footnote\@gobble
339 | \let\thanks\@gobble
340 | \def\footnotemark{}
341 | \def\and{and }%
342 | \def\,{ }
343 | \def\\{ }
344 | }{%
345 | \string\contitem
346 | \string\title{\@title}%
347 | \string\author{\@authorsfortoc}%
348 | \string\page{\@ArticleNo:\thepage--\@ArticleNo:\number\numexpr\getpagerefnumber{TotPages}}}%
349 | \closeout\tocfile
350 | \egroup
351 | \par}
352 | \renewcommand\tableofcontents{%
353 | \section*{\contentsname}%
354 | \@starttoc{toc}}
355 | \setcounter{secnumdepth}{4}
356 | \renewcommand\section{\@startsection {section}{1}{\z@}%
357 | {-3.5ex \@plus -1ex \@minus -.2ex}%
358 | {2.3ex \@plus.2ex}%
359 | {\sffamily\Large\bfseries\raggedright}}
360 | \renewcommand\subsection{\@startsection{subsection}{2}{\z@}%
361 | {-3.25ex\@plus -1ex \@minus -.2ex}%
362 | {1.5ex \@plus .2ex}%
363 | {\sffamily\Large\bfseries\raggedright}}
364 | \renewcommand\subsubsection{\@startsection{subsubsection}{3}{\z@}%
365 | {-3.25ex\@plus -1ex \@minus -.2ex}%
366 | {1.5ex \@plus .2ex}%
367 | {\sffamily\Large\bfseries\raggedright}}
368 | \renewcommand\paragraph{\@startsection{paragraph}{4}{\z@}%
369 | {-3.25ex \@plus-1ex \@minus-.2ex}%
370 | {1.5ex \@plus .2ex}%
371 | {\sffamily\large\bfseries\raggedright}}
372 | \renewcommand\subparagraph{\@startsection{subparagraph}{5}{\z@}%
373 | {3.25ex \@plus1ex \@minus .2ex}%
374 | {-1em}%
375 | {\sffamily\normalsize\bfseries}}
376 | \newcommand{\proofsubparagraph}{\@startsection{subparagraph}{5}{\z@}%
377 | {3.25ex \@plus1ex \@minus .2ex}%
378 | {-1em}%
379 | {\color{lipicsGray}\sffamily\normalsize\bfseries}}
380 | \setlength\leftmargini \parindent
381 | \setlength\leftmarginii {1.2em}
382 | \setlength\leftmarginiii{1.2em}
383 | \setlength\leftmarginiv {1.2em}
384 | \setlength\leftmarginv {1.2em}
385 | \setlength\leftmarginvi {1.2em}
386 | \renewcommand\labelenumi{%
387 | \textcolor{lipicsGray}{\sffamily\bfseries\upshape\mathversion{bold}\theenumi.}}
388 | \renewcommand\labelenumii{%
389 | \textcolor{lipicsGray}{\sffamily\bfseries\upshape\mathversion{bold}\theenumii.}}
390 | \renewcommand\labelenumiii{%
391 | \textcolor{lipicsGray}{\sffamily\bfseries\upshape\mathversion{bold}\theenumiii.}}
392 | \renewcommand\labelenumiv{%
393 | \textcolor{lipicsGray}{\sffamily\bfseries\upshape\mathversion{bold}\theenumiv.}}
394 | \renewcommand\labelitemi{%
395 | \textcolor{lipicsBulletGray}{\ifnum\@listdepth=\@ne
396 | \rule{0.67em}{0.33em}%
397 | \else
398 | \rule{0.45em}{0.225em}%
399 | \fi}}
400 | \renewcommand\labelitemii{%
401 | \textcolor{lipicsBulletGray}{\rule{0.45em}{0.225em}}}
402 | \renewcommand\labelitemiii{%
403 | \textcolor{lipicsBulletGray}{\sffamily\bfseries\textasteriskcentered}}
404 | \renewcommand\labelitemiv{%
405 | \textcolor{lipicsBulletGray}{\sffamily\bfseries\textperiodcentered}}
406 | \renewenvironment{description}
407 | {\list{}{\advance\partopsep\topsep\topsep\z@\@plus\p@
408 | \labelwidth\z@ \itemindent-\leftmargin
409 | \let\makelabel\descriptionlabel}}
410 | {\ifnum\@listdepth=\@ne\global\@nobotseplisttrue\fi\endlist}
411 | \renewcommand*\descriptionlabel[1]{%
412 | \hspace\labelsep\textcolor{lipicsGray}{\sffamily\bfseries\mathversion{bold}#1}}
413 | \def\topmattervskip{0.7}
414 | \renewenvironment{abstract}{%
415 | \vskip\topmattervskip\bigskipamount
416 | \noindent
417 | \rlap{\color{lipicsLineGray}\vrule\@width\textwidth\@height1\p@}%
418 | \hspace*{7mm}\fboxsep1.5mm\colorbox[rgb]{1,1,1}{\raisebox{-0.4ex}{%
419 | \large\selectfont\sffamily\bfseries\abstractname}}%
420 | \vskip3\p@
421 | \fontsize{9}{12}\selectfont
422 | \noindent\ignorespaces}
423 | {\vskip\topmattervskip\baselineskip\noindent
424 | \subjclassHeading
425 | \ifx\@ccsdescString\@empty
426 | \textcolor{red}{Author: Please fill in 1 or more \string\ccsdesc\space macro}%
427 | \else
428 | \@ccsdescString
429 | \fi
430 | \vskip\topmattervskip\baselineskip
431 | \noindent\keywordsHeading
432 | \ifx\@keywords\@empty
433 | \textcolor{red}{Author: Please fill in \string\keywords\space macro}%
434 | \else
435 | \@keywords
436 | \fi
437 | \ifx\@hideLIPIcs\@undefined
438 | \ifx\@DOIPrefix\@empty\else
439 | \vskip\topmattervskip\baselineskip\noindent
440 | \doiHeading\href{https://doi.org/\@lipicsdoi}{\@lipicsdoi}%
441 | \fi
442 | \fi
443 | \ifx\@category\@empty\else
444 | \vskip\topmattervskip\baselineskip\noindent
445 | \categoryHeading\@category
446 | \fi
447 | \ifx\@relatedversion\@empty\else
448 | \vskip\topmattervskip\baselineskip\noindent
449 | \relatedversionHeading\ifx\authoranonymous\relax\textcolor{red}{Anonymous related version(s)}\else\@relatedversion\fi
450 | \fi
451 | \ifx\@supplement\@empty\else
452 | \vskip\topmattervskip\baselineskip\noindent
453 | \supplementHeading\ifx\authoranonymous\relax\textcolor{red}{Anonymous supplemenatary material}\else\@supplement\fi
454 | \fi
455 | \ifx\@funding\@empty\else
456 | \vskip\topmattervskip\baselineskip\noindent
457 | \fundingHeading\ifx\authoranonymous\relax\textcolor{red}{Anonymous funding}\else\@funding\fi
458 | \fi
459 | \ifx\@acknowledgements\@empty\else
460 | \vskip\topmattervskip\baselineskip\noindent
461 | \acknowledgementsHeading\ifx\authoranonymous\relax\textcolor{red}{Anonymous acknowledgements} \else\@acknowledgements\fi
462 | \fi
463 | \protected@write\@auxout{}{\string\gdef\string\@pageNumberEndAbstract{\thepage}}%
464 | }% end abstract
465 | \renewenvironment{thebibliography}[1]
466 | {\if@noskipsec \leavevmode \fi
467 | \par
468 | \@tempskipa-3.5ex \@plus -1ex \@minus -.2ex\relax
469 | \@afterindenttrue
470 | \@tempskipa -\@tempskipa \@afterindentfalse
471 | \if@nobreak
472 | \everypar{}%
473 | \else
474 | \addpenalty\@secpenalty\addvspace\@tempskipa
475 | \fi
476 | \noindent
477 | \rlap{\color{lipicsLineGray}\vrule\@width\textwidth\@height1\p@}%
478 | \hspace*{7mm}\fboxsep1.5mm\colorbox[rgb]{1,1,1}{\raisebox{-0.4ex}{%
479 | \normalsize\sffamily\bfseries\refname}}%
480 | \@xsect{1ex \@plus.2ex}%
481 | \list{\@biblabel{\@arabic\c@enumiv}}%
482 | {\leftmargin8.5mm
483 | \labelsep\leftmargin
484 | \settowidth\labelwidth{\@biblabel{#1}}%
485 | \advance\labelsep-\labelwidth
486 | \usecounter{enumiv}%
487 | \let\p@enumiv\@empty
488 | \renewcommand\theenumiv{\@arabic\c@enumiv}}%
489 | \fontsize{9}{12}\selectfont
490 | \sloppy
491 | \clubpenalty4000
492 | \@clubpenalty \clubpenalty
493 | \widowpenalty4000%
494 | \sfcode`\.\@m\protected@write\@auxout{}{\string\gdef\string\@pageNumberStartBibliography{\thepage}}}
495 | {\def\@noitemerr
496 | {\@latex@warning{Empty `thebibliography' environment}}%
497 | \protected@write\@auxout{}{\string\gdef\string\@pageNumberEndBibliography{\thepage}}%
498 | \endlist}
499 | \g@addto@macro\appendix{\immediate\write\@auxout{\string\gdef\string\@pageNumberStartAppendix{\thepage}}}%
500 | \renewcommand\footnoterule{%
501 | \kern-8\p@
502 | {\color{lipicsBulletGray}\hrule\@width40mm\@height1\p@}%
503 | \kern6.6\p@}
504 | \renewcommand\@makefntext[1]{%
505 | \parindent\z@\hangindent1em
506 | \leavevmode
507 | \hb@xt@1em{\@makefnmark\hss}#1}
508 | \usepackage{microtype}
509 | \usepackage[utf8]{inputenc}
510 | \IfFileExists{glyphtounicode.tex}{
511 | \input glyphtounicode
512 | \pdfgentounicode=1
513 | }{}%
514 | \IfFileExists{lmodern.sty}{\RequirePackage{lmodern}}{}
515 | \IfFileExists{fontawesome5.sty}{%
516 | \RequirePackage{fontawesome5}%
517 | \IfFileExists{orcid.pdf}{%
518 | \def\orcidsymbol{\includegraphics[height=9\p@]{orcid}}
519 | }{
520 | \def\orcidsymbol{\textcolor{lipicsGray}{\fontsize{9}{12}\sffamily\bfseries \faOrcid}}%
521 | }
522 | \def\mailsymbol{\textcolor{lipicsGray}{\fontsize{9}{12}\sffamily\bfseries \faIcon[regular]{envelope}}}%
523 | \def\homesymbol{\textcolor{lipicsGray}{\fontsize{9}{12}\sffamily\bfseries \faHome}}%
524 | }{%
525 | \ClassWarning{Package fontawesome5 not installed}{Please install package fontawesome5}
526 | \def\orcidsymbol{\textcolor{lipicsGray}{\fontsize{9}{12}\sffamily\bfseries ORCID}}
527 | \def\mailsymbol{\textcolor{lipicsGray}{\fontsize{9}{12}\sffamily\bfseries @}}%
528 | \def\homesymbol{\textcolor{lipicsGray}{\fontsize{9}{12}\sffamily\bfseries H}}%
529 | }%
530 | \RequirePackage[T1]{fontenc}
531 | \RequirePackage{textcomp}
532 | \RequirePackage[mathscr]{eucal}
533 | \RequirePackage{amssymb}
534 | \PassOptionsToPackage{retainmissing}{MnSymbol}
535 | \AtBeginDocument{\@ifpackageloaded{MnSymbol}%
536 | {\expandafter\let\csname ver@amssymb.sty\endcsname\relax
537 | \let\complement\@undefined
538 | \RequirePackage{amssymb}}{}}
539 | \RequirePackage{soul}
540 | \sodef\textsolittle{}{.12em}{.5em\@plus.08em\@minus.06em}%
541 | {.4em\@plus.275em\@minus.183em}
542 | \RequirePackage{color} %kept for backward compatibility
543 | \AtBeginDocument{
544 | \@ifpackageloaded{xcolor}{
545 | }{
546 | \RequirePackage{xcolor}
547 | }
548 | \definecolor{darkgray}{rgb}{0.31,0.31,0.33}
549 | \definecolor[named]{lipicsGray}{rgb}{0.31,0.31,0.33}
550 | \definecolor[named]{lipicsBulletGray}{rgb}{0.60,0.60,0.61}
551 | \definecolor[named]{lipicsLineGray}{rgb}{0.51,0.50,0.52}
552 | \definecolor[named]{lipicsLightGray}{rgb}{0.85,0.85,0.86}
553 | \definecolor[named]{lipicsYellow}{rgb}{0.99,0.78,0.07}
554 | }
555 | \RequirePackage{babel}
556 | \RequirePackage[tbtags,fleqn]{amsmath}
557 | \AtBeginDocument{
558 | \@ifpackageloaded{enumitem}{\ClassWarning{Package 'enuitem' incompatible}{Don't use package 'enumitem'; Package enumerate preloaded!}}{}
559 | \@ifpackageloaded{paralist}{\ClassWarning{Package 'paralist' incompatible}{Don't use package 'paralist'; Package enumerate preloaded!}}{}
560 | }
561 | \RequirePackage{enumerate}
562 | \def\@enum@{\list{\textcolor{lipicsGray}{\sffamily\bfseries\upshape\mathversion{bold}\csname label\@enumctr\endcsname}}%
563 | {\advance\partopsep\topsep
564 | \topsep\z@\@plus\p@
565 | \usecounter{\@enumctr}\def\makelabel##1{\hss\llap{##1}}}}
566 | \def\romanenumerate{\enumerate[(i)]}
567 | \let\endromanenumerate\endenumerate
568 | \def\alphaenumerate{\enumerate[(a)]}
569 | \let\endalphaenumerate\endenumerate
570 | \def\bracketenumerate{\enumerate[(1)]}
571 | \let\endbracketenumerate\endenumerate
572 | \RequirePackage{graphicx}
573 | \RequirePackage{array}
574 | \let\@classzold\@classz
575 | \def\@classz{%
576 | \expandafter\ifx\d@llarbegin\begingroup
577 | \toks \count@ =
578 | \expandafter{\expandafter\small\the\toks\count@}%
579 | \fi
580 | \@classzold}
581 | \RequirePackage{multirow}
582 | \RequirePackage{tabularx}
583 | \RequirePackage[online]{threeparttable}
584 | \def\TPTtagStyle#1{#1)}
585 | \def\tablenotes{\small\TPT@defaults
586 | \@ifnextchar[\TPT@setuptnotes\TPTdoTablenotes} % ]
587 | \RequirePackage{listings}
588 | \lstset{basicstyle=\small\ttfamily,%
589 | backgroundcolor=\color{lipicsLightGray},%
590 | frame=single,framerule=0pt,xleftmargin=\fboxsep,xrightmargin=\fboxsep}
591 | \RequirePackage[left,mathlines]{lineno}
592 | \linenumbers
593 | \renewcommand\linenumberfont{\normalfont\tiny\sffamily}
594 | %%%% patch to cope with amsmath
595 | %%%% http://phaseportrait.blogspot.de/2007/08/lineno-and-amsmath-compatibility.html
596 | \newcommand*\patchAmsMathEnvironmentForLineno[1]{%
597 | \expandafter\let\csname old#1\expandafter\endcsname\csname #1\endcsname
598 | \expandafter\let\csname oldend#1\expandafter\endcsname\csname end#1\endcsname
599 | \renewenvironment{#1}%
600 | {\linenomath\csname old#1\endcsname}%
601 | {\csname oldend#1\endcsname\endlinenomath}}%
602 | \newcommand*\patchBothAmsMathEnvironmentsForLineno[1]{%
603 | \patchAmsMathEnvironmentForLineno{#1}%
604 | \patchAmsMathEnvironmentForLineno{#1*}}%
605 | \AtBeginDocument{%
606 | \patchBothAmsMathEnvironmentsForLineno{equation}%
607 | \patchBothAmsMathEnvironmentsForLineno{align}%
608 | \patchBothAmsMathEnvironmentsForLineno{flalign}%
609 | \patchBothAmsMathEnvironmentsForLineno{alignat}%
610 | \patchBothAmsMathEnvironmentsForLineno{gather}%
611 | \patchBothAmsMathEnvironmentsForLineno{multline}}
612 | \let\usehyperxmp\@empty%
613 | \ifx\pdfa\relax%
614 | \IfFileExists{hyperxmp.sty}{%
615 | \RequirePackage{hyperxmp}%
616 | \@ifpackagelater{hyperxmp}{2019/04/05}{%
617 | \let\usehyperxmp\relax%
618 | }{%
619 | \ClassWarning{Package hyperxmp outdated}{You are using an outdated version of the package hyperxmp. Please update!}%
620 | }}{}%
621 | \fi%
622 | \IfFileExists{totpages.sty}{
623 | \RequirePackage{totpages}
624 | }{
625 | \ClassWarning{Package totpages not installed}{Please install package totpages}
626 | \newcounter{TotPages}
627 | \setcounter{TotPages}{99}
628 | }
629 | \ifx\usehyperxmp\relax%
630 | \RequirePackage[pdfa,unicode]{hyperref}%
631 | \else%
632 | \RequirePackage[unicode]{hyperref}%
633 | \fi%
634 | \let\C\relax%
635 | \let\G\relax%
636 | \let\F\relax%
637 | \let\U\relax%
638 | \pdfstringdefDisableCommands{%
639 | \let\thanks\@gobble%
640 | \let\footnote\@gobble%
641 | \def\footnotemark{}%
642 | \def\cs#1{\textbackslash #1}%
643 | \let\normalfont\@empty%
644 | \let\scshape\@empty%
645 | \def\and{and }%
646 | \def\,{ }%
647 | \def\textrightarrow{ -> }%
648 | \let\mathsf\@empty%
649 | }%
650 | \hypersetup{
651 | breaklinks=true,
652 | pdfencoding=unicode,
653 | bookmarksnumbered,
654 | pdfborder={0 0 0},
655 | pdfauthor={ }
656 | }%
657 | \AtBeginDocument{
658 | \ifx\usehyperxmp\relax
659 | \hypersetup{
660 | pdftitle={\@title},
661 | pdfauthor={\ifx\authoranonymous\relax Anonymous author(s) \else \@authorsforpdf \fi},
662 | pdfkeywords={\@keywords},
663 | pdfproducer={LaTeX with lipics-v2021.cls},
664 | pdfsubject={LIPIcs, Vol.\@SeriesVolume, \@EventShortTitle},
665 | pdfcopyright = { Copyright (C) \ifx\authoranonymous\relax Anonymous author(s) \else \@copyrightholder; \fi licensed under Creative Commons License CC-BY 4.0},
666 | pdflang={en},
667 | pdfmetalang={en},
668 | pdfpublisher={Schloss Dagstuhl -- Leibniz-Zentrum für Informatik, Dagstuhl Publishing, Germany},
669 | pdflicenseurl={https://creativecommons.org/licenses/by/4.0/},
670 | pdfpubtype={LIPIcs},
671 | pdfvolumenum={\@SeriesVolume},
672 | pdfpagerange={\@ArticleNo:\thepage-\@ArticleNo:\theTotPages},
673 | pdfdoi={\@lipicsdoi},
674 | pdfapart=3,
675 | pdfaconformance=B
676 | }
677 | \else%
678 | \hypersetup{
679 | pdftitle={\@title},
680 | pdfauthor={\ifx\authoranonymous\relax Anonymous author(s) \else \@authorsforpdf \fi},
681 | pdfkeywords={\@keywords},
682 | pdfcreator={LaTeX with lipics-v2021.cls},
683 | pdfsubject={LIPIcs, Vol.\@SeriesVolume, \@EventShortTitle; Copyright (C) \ifx\authoranonymous\relax Anonymous author(s) \else \@copyrightholder; \fi licensed under Creative Commons License CC-BY 4.0}
684 | }%
685 | \fi %
686 | }
687 | \ifx\usehyperxmp\relax
688 | \pdfobjcompresslevel=0
689 | \pdfinclusioncopyfonts=1
690 | \IfFileExists{colorprofiles.tex}{
691 | \RequirePackage{colorprofiles}%
692 | \IfFileExists{sRGB.icc}{
693 | \immediate\pdfobj stream attr{/N 3} file{sRGB.icc}
694 | \pdfcatalog{%
695 | /OutputIntents [
696 | <<
697 | /Type /OutputIntent
698 | /S /GTS_PDFA1
699 | /DestOutputProfile \the\pdflastobj\space 0 R
700 | /OutputConditionIdentifier (sRGB)
701 | /Info (sRGB)
702 | >>
703 | ]
704 | }}{}
705 | }{\ClassWarning{Package colorprofiles not installed}{Please install package colorprofiles}}
706 | \fi
707 | \RequirePackage[labelsep=space,singlelinecheck=false,%
708 | font={up,small},labelfont={sf,bf},%
709 | listof=false]{caption}%"listof" instead of "list" for backward compatibility
710 | \@ifpackagelater{hyperref}{2009/12/09}
711 | {\captionsetup{compatibility=false}}%cf. http://groups.google.de/group/comp.text.tex/browse_thread/thread/db9310eb540fbbd8/42e30f3b7b3aa17a?lnk=raot
712 | {}
713 | \DeclareCaptionLabelFormat{boxed}{%
714 | \kern0.05em{\color[rgb]{0.99,0.78,0.07}\rule{0.73em}{0.73em}}%
715 | \hspace*{0.67em}\bothIfFirst{#1}{~}#2}
716 | \captionsetup{labelformat=boxed}
717 | \captionsetup[table]{position=top}
718 | \RequirePackage[figuresright]{rotating}
719 | \caption@AtBeginDocument{\@ifpackageloaded{subfig}{\ClassError{lipics}{%
720 | Do not load the subfig package}{The more recent subcaption package is already loaded}}{}}
721 | \RequirePackage{subcaption}
722 | \def\titlerunning#1{\gdef\@titlerunning{{\let\footnote\@gobble\markboth{#1}{#1}}}}
723 | \def\authorrunning#1{%
724 | \gdef\@authorrunning{\markright{\ifx\authoranonymous\relax\textcolor{red}{Anonymous author(s)} \else\if!#1!\textcolor{red}{Author: Please fill in the \string\authorrunning\space macro}\else#1\fi\fi}}}
725 | \titlerunning{\@title}
726 | \authorrunning{\textcolor{red}{Author: Please use the \string\authorrunning\space macro}}
727 | \def\EventLongTitle#1{\gdef\@EventLongTitle{#1}}
728 | \EventLongTitle{}
729 | \def\EventShortTitle#1{\gdef\@EventShortTitle{#1}}
730 | \EventShortTitle{}
731 | \def\EventEditors#1{\gdef\@EventEditors{#1}}
732 | \EventEditors{}
733 | \def\EventNoEds#1{\gdef\@EventNoEds{#1}\xdef\@Eds{Editor\ifnum#1>1s\fi}}
734 | \EventNoEds{1}
735 | \def\EventLogo#1{\gdef\@EventLogo{#1}}
736 | \EventLogo{}
737 | \def\EventAcronym#1{\gdef\@EventAcronym{#1}}
738 | \EventAcronym{}
739 | \def\EventYear#1{\gdef\@EventYear{#1}}
740 | \EventYear{}
741 | \def\EventDate#1{\gdef\@EventDate{#1}}
742 | \EventDate{}
743 | \def\EventLocation#1{\gdef\@EventLocation{#1}}
744 | \EventLocation{}
745 | \def\SeriesVolume#1{\gdef\@SeriesVolume{#1}}
746 | \SeriesVolume{}
747 | \def\ArticleNo#1{\gdef\@ArticleNo{#1}}
748 | \ArticleNo{}
749 | \def\DOIPrefix#1{\gdef\@DOIPrefix{#1}}
750 | \DOIPrefix{10.4230/LIPIcs}
751 | \def\@lipicsdoi{\@DOIPrefix.\@EventAcronym.\@EventYear.\@ArticleNo}
752 | \def\and{\newline}
753 | \let\orig@author\@author
754 | \let\@authorsfortoc\@empty
755 | \let\@authorsforpdf\@empty
756 | \newcount\c@author
757 | \newcounter{currentauthor}
758 | \def\authorcolumnsMin{6}
759 | \def\@authornum{0}
760 | \def\author#1#2#3#4#5{%
761 | \ifx\@author\orig@author\let\@author\@empty\fi
762 | \g@addto@macro\@author{%
763 | \noexpandarg\StrBehind{#2}{\and \url}[\homepageTemp]\IfSubStr{#2}{\and \url}{\StrBefore{#2}{\and \url}[\affiliation]}{\def\affiliation{#2}}%
764 | \expandarg\exploregroups\StrRemoveBraces{\homepageTemp}[\homepage]%
765 | \ifx\authorcolumns\relax
766 | \ifnum\c@author>\authorcolumnsMin
767 | \stepcounter{currentauthor}
768 | \ifodd\value{currentauthor}
769 | \begin{minipage}[t]{\textwidth}
770 | \begin{minipage}[t]{0.49\textwidth}
771 | \else
772 | \hfill \begin{minipage}[t]{0.49\textwidth}
773 | \fi
774 | \else
775 | \ClassWarning{Option 'authorcolumns' only applicable for > 6 authors}{Option 'authorcolumns' only applicable for >6 authors!}
776 | \addvspace{0.5\baselineskip}
777 | \fi
778 | \else
779 | \addvspace{0.5\baselineskip}
780 | \fi
781 | {\Large\bfseries
782 | \if!#1!
783 | \textcolor{red}{Author: Please enter author name}%
784 | \else
785 | \ifx\authoranonymous\relax
786 | \textcolor{red}{Anonymous author}
787 | \else
788 | #1\,%
789 | \ifx\compactauthor\relax\if!#3!\else{\,\href{mailto:#3}{\mailsymbol}}\fi%
790 | \ifx\homepage\@empty\else{\,\href{\homepage}{\homesymbol}}\fi\fi%
791 | \if!#4!\else{\,\href{#4}{\orcidsymbol}}\fi%
792 | \if!#5!\else
793 | \ifx\@funding\@empty
794 | \expandafter\g@addto@macro\expandafter\@funding{\textit{\expandafter{\let\footnote\@gobble #1}}:\space{#5}}
795 | \else
796 | \expandafter\g@addto@macro\expandafter\@funding{\\\textit{\expandafter{\let\footnote\@gobble #1}}:\space{#5}}
797 | \fi
798 | \fi
799 | \fi
800 | \fi
801 | }
802 | {\small
803 | \if!#2!\textcolor{red}{Author: Please enter affiliation as second parameter of the author macro}\else{\\* \ifx\authoranonymous\relax\textcolor{red}{Anonymous affiliation}\else\ifx\compactauthor\relax \affiliation \else#2\fi\fi}\fi
804 | \ifx\compactauthor\relax\else\if!#3!\else{\ifx\authoranonymous\relax\else\\*\href{mailto:#3}{#3}\fi}\fi\fi
805 | }\par
806 | \ifx\authorcolumns\relax
807 | \ifnum\c@author>\authorcolumnsMin
808 | \end{minipage}
809 | \ifnum\c@author=\value{currentauthor}
810 | \end{minipage}
811 | \else
812 | \ifodd\value{currentauthor}
813 | \else
814 | \end{minipage}%
815 | \medskip
816 | \fi
817 | \fi
818 | \fi
819 | \fi}%
820 | \global\advance\c@author\@ne
821 | \protected@write\@auxout{}{\string\gdef\string\@authornum{\the\c@author}}
822 | \ifnum\c@author=\@ne
823 | \gdef\@authorsfortoc{#1}%
824 | \gdef\@authorsforpdf{#1}
825 | \else
826 | \expandafter\g@addto@macro\expandafter\@authorsforpdf\expandafter{, #1}
827 | \expandafter\g@addto@macro\expandafter\@authorsfortoc\expandafter{\expandafter\csname\the\c@author authand\endcsname#1}%
828 | \@namedef{\the\c@author authand}{,\space}%
829 | \AtBeginDocument{%
830 | \expandafter\ifnum\@authornum=2
831 | \@namedef{2authand}{\space and\space}%
832 | \else
833 | \@namedef{\@authornum authand}{,\space and\space}%
834 | \fi}
835 | \fi}
836 | \newcommand*\affil[2][]{%
837 | \ClassError{lipics}
838 | {\string\affil\space deprecated: Please enter affiliation as second parameter of the author macro}
839 | {Since 2017, \string\affil\space is obsolete in lipics.}}
840 | \newcommand*\Copyright[1]{%
841 | \def\@copyrightholder{#1}
842 | \def\@Copyright{%
843 | \setbox\@tempboxa\hbox{\IfFileExists{cc-by.pdf}{\includegraphics[height=14\p@,clip]{cc-by}}{\includegraphics[height=14\p@, width=40pt]{example-image-plain}}}%
844 | \@rightskip\@flushglue \rightskip\@rightskip
845 | \hangindent\dimexpr\wd\@tempboxa+0.5em\relax
846 | \href{https://creativecommons.org/licenses/by/4.0/}%
847 | {\smash{\lower\baselineskip\hbox{\unhcopy\@tempboxa}}}\enskip
848 | \textcopyright\ %
849 | \ifx!#1!\textcolor{red}{Author: Please fill in the \string\Copyright\space macro}\else\ifx\authoranonymous\relax\textcolor{red}{Anonymous author(s)}\else#1\fi\fi
850 | ;\\%
851 | licensed under Creative Commons License CC-BY 4.0\ifx!#1!\\\null\fi\par}}
852 | \Copyright{\textcolor{red}{Author: Please provide a copyright holder}}
853 | \let\@copyrightholder\@empty
854 | \def\hideLIPIcs{\let\@hideLIPIcs\relax}
855 | \usepackage{xstring}
856 | \def\keywords#1{\def\@keywords{#1}}
857 | \let\@keywords\@empty
858 | \def\keywordsHeading{%
859 | \textcolor{lipicsGray}{\fontsize{9}{12}\sffamily\bfseries
860 | Keywords and phrases\enskip}}
861 | \RequirePackage{comment}
862 | \excludecomment{CCSXML}
863 | % inspired by https://tex.stackexchange.com/questions/12810/how-do-i-split-a-string
864 | \global\newcommand\ccsdesc[2][100]{\@ccsdesc#1~#2~~\relax}
865 | \let\orig@ccsdesc\@ccsdesc
866 | \let\@ccsdesc\@empty
867 | \let\@ccsdescString\@empty
868 | \gdef\@ccsdesc#1~#2~#3~{
869 | \ifx\@ccsdesc\orig@ccsdesc\let\@ccsdesc\@empty\fi
870 | \ifx!#3!
871 | \ifx\@ccsdescString\@empty
872 | \g@addto@macro\@ccsdescString{{#2}}
873 | \else
874 | \g@addto@macro\@ccsdescString{; {#2}}
875 | \fi
876 | \else
877 | \ifx\@ccsdescString\@empty
878 | \g@addto@macro\@ccsdescString{{#2} $\rightarrow$ {#3}}
879 | \else
880 | \g@addto@macro\@ccsdescString{; {#2} $\rightarrow$ {#3}}
881 | \fi
882 | \fi
883 | \ccsdescEnd
884 | }
885 | \def\ccsdescEnd#1\relax{}
886 | \def\subjclass#1{
887 | \ClassError{lipics}
888 | {\string\subjclass\space deprecated: Please enter subject classification in 1 or more ccsdesc macros}
889 | {Since 2019, \string\subjclass\space is obsolete in lipics.}}
890 | \let\@subjclass\@empty
891 | \def\subjclassHeading{%
892 | \textcolor{lipicsGray}{\fontsize{9}{12}\sffamily\bfseries
893 | 2012 ACM Subject Classification\enskip}}
894 | \def\doiHeading{%
895 | \textcolor{lipicsGray}{\fontsize{9}{12}\sffamily\bfseries
896 | Digital Object Identifier\enskip}}
897 | \def\category#1{\def\@category{#1}}
898 | \let\@category\@empty
899 | \def\categoryHeading{%
900 | \textcolor{lipicsGray}{\fontsize{9}{12}\sffamily\bfseries
901 | Category\enskip}}
902 | \def\relatedversion#1{\def\@relatedversion{#1}}
903 | \let\@relatedversion\@empty
904 | \define@key{relatedversiondetails}{linktext}{\def\relatedversiondetails@linktext{#1}}
905 | \define@key{relatedversiondetails}{cite}{\def\relatedversiondetails@cite{#1}}
906 | \newcommand*\addtorelatedversionmacro[2]{%
907 | \ifx\@relatedversion\@empty%
908 | \g@addto@macro\@relatedversion{#1}%
909 | \else%
910 | \g@addto@macro\@relatedversion{\\#1}%
911 | \fi%
912 | }%
913 | \newcommand{\relatedversiondetails}[3][]{%
914 | \begingroup%
915 | \let\relatedversiondetails@linktext\@empty
916 | \let\relatedversiondetails@cite\@empty
917 | \setkeys{relatedversiondetails}{#1}%
918 | \ifx\relatedversiondetails@linktext\@empty%
919 | \protected@edef\tmp{\textit{#2}:\space{\url{#3}}}%
920 | \else%
921 | \protected@edef\tmp{\textit{#2}:\space{\href{#3}{\texttt{\relatedversiondetails@linktext}}}}%
922 | \fi%
923 | \ifx\relatedversiondetails@cite\@empty%
924 | \else%
925 | \protected@edef\tmp{\tmp\space\cite{\relatedversiondetails@cite}}%
926 | \fi%
927 | \expandafter\addtorelatedversionmacro\expandafter{\tmp}{#1}%
928 | \endgroup%
929 | }%
930 | \def\relatedversionHeading{%
931 | \textcolor{lipicsGray}{\fontsize{9}{12}\sffamily\bfseries
932 | Related Version\enskip}}
933 | \def\supplement#1{\def\@supplement{#1}}
934 | \let\@supplement\@empty
935 | \define@key{supplementdetails}{linktext}{\def\supplementdetails@linktext{#1}}
936 | \define@key{supplementdetails}{cite}{\def\supplementdetails@cite{#1}}
937 | \define@key{supplementdetails}{subcategory}{\def\supplementdetails@subcategory{#1}}
938 | \define@key{supplementdetails}{swhid}{\def\supplementdetails@swhid{#1}}
939 | \newcommand*\addtosupplementmacro[2]{%
940 | \ifx\@supplement\@empty%
941 | \g@addto@macro\@supplement{#1}%
942 | \else%
943 | \g@addto@macro\@supplement{\\#1}%
944 | \fi%
945 | }%
946 | \newcommand{\supplementdetails}[3][]{%
947 | \begingroup%
948 | \let\supplementdetails@linktext\@empty
949 | \let\supplementdetails@cite\@empty
950 | \let\supplementdetails@subcategory\@empty
951 | \let\supplementdetails@swhid\@empty
952 | \setkeys{supplementdetails}{#1}%
953 | \ifx\supplementdetails@subcategory\@empty%
954 | \protected@edef\tmp{\textit{#2}}
955 | \else
956 | \protected@edef\tmp{\textit{#2\,\,(\supplementdetails@subcategory)}}%
957 | \fi
958 | \ifx\supplementdetails@linktext\@empty%
959 | \protected@edef\tmp{\tmp:\space{\url{#3}}}%
960 | \else%
961 | \protected@edef\tmp{\tmp:\space{\href{#3}{\texttt{\supplementdetails@linktext}}}}%
962 | \fi%
963 | \ifx\supplementdetails@cite\@empty%
964 | \else%
965 | \protected@edef\tmp{\tmp\space\cite{\supplementdetails@cite}}%
966 | \fi
967 | \ifx\supplementdetails@swhid\@empty%
968 | \else%
969 | \protected@edef\tmp{\tmp \\ \hspace*{1.2em} archived at \href{https://archive.softwareheritage.org/\supplementdetails@swhid}{\texttt{\supplementdetails@swhid}}}%
970 | \fi%
971 | \expandafter\addtosupplementmacro\expandafter{\tmp}{#1}%
972 | \endgroup%
973 | }%
974 | \def\supplementHeading{%
975 | \textcolor{lipicsGray}{\fontsize{9}{12}\sffamily\bfseries
976 | Supplementary Material\enskip}}
977 | \newcommand\flag[2][0.9cm]{%
978 | \leavevmode\marginpar{%
979 | \raisebox{\dimexpr-\totalheight+\ht\strutbox\relax}%
980 | [\dimexpr\ht\strutbox+3mm][\dp\strutbox]{\expandafter\includegraphics[width=#1]{#2}}%
981 | }}
982 | \def\funding#1{\def\@funding{#1}}
983 | \let\@funding\@empty
984 | \def\fundingHeading{%
985 | \textcolor{lipicsGray}{\fontsize{9}{12}\sffamily\bfseries
986 | Funding\enskip}}
987 | \def\acknowledgements#1{\def\@acknowledgements{#1}}
988 | \let\@acknowledgements\@empty
989 | \def\acknowledgementsHeading{%
990 | \textcolor{lipicsGray}{\fontsize{9}{12}\sffamily\bfseries
991 | Acknowledgements\enskip}}
992 | \RequirePackage{amsthm}
993 | \ifx\usethmrestate\relax
994 | \RequirePackage{thm-restate}
995 | \fi
996 | \thm@headfont{%
997 | \textcolor{lipicsGray}{$\blacktriangleright$}\nobreakspace\sffamily\bfseries}
998 | \def\th@remark{%
999 | \thm@headfont{%
1000 | \textcolor{lipicsGray}{$\blacktriangleright$}\nobreakspace\sffamily}%
1001 | \normalfont % body font
1002 | \thm@preskip\topsep \divide\thm@preskip\tw@
1003 | \thm@postskip\thm@preskip
1004 | }
1005 | \def\@endtheorem{\endtrivlist}%\@endpefalse
1006 | \renewcommand\qedsymbol{\textcolor{lipicsGray}{\ensuremath{\blacktriangleleft}}}
1007 | \renewenvironment{proof}[1][\proofname]{\par
1008 | \pushQED{\qed}%
1009 | \normalfont \topsep6\p@\@plus6\p@\relax
1010 | \trivlist
1011 | \item[\hskip\labelsep
1012 | \color{lipicsGray}\sffamily\bfseries
1013 | #1\@addpunct{.}]\ignorespaces
1014 | }{%
1015 | \popQED\endtrivlist%\@endpefalse
1016 | }
1017 | \newcommand{\claimqedhere}{\renewcommand\qedsymbol{\textcolor{lipicsGray}{\ensuremath{\vartriangleleft}}}%
1018 | \qedhere%
1019 | \renewcommand\qedsymbol{\textcolor{lipicsGray}{\ensuremath{\blacktriangleleft}}}}
1020 | \newenvironment{claimproof}[1][\proofname]{
1021 | \pushQED{\qed}%
1022 | \normalfont \topsep6\p@\@plus6\p@\relax
1023 | \trivlist
1024 | \item[\hskip\labelsep
1025 | \color{lipicsGray}\sffamily
1026 | #1\@addpunct{.}]\ignorespaces
1027 | }{%
1028 | \renewcommand\qedsymbol{\textcolor{lipicsGray}{\ensuremath{\vartriangleleft}}}
1029 | \popQED\endtrivlist%\@endpefalse
1030 | \renewcommand\qedsymbol{\textcolor{lipicsGray}{\ensuremath{\blacktriangleleft}}}
1031 | }
1032 | % inspired by qed of amsthm class
1033 | \DeclareRobustCommand{\lipicsEnd}{%
1034 | \leavevmode\unskip\penalty9999 \hbox{}\nobreak\hfill
1035 | \quad\hbox{$\lrcorner$}%
1036 | }
1037 | \AtBeginDocument{
1038 | \@ifpackageloaded{algorithm2e}{
1039 | \@ifpackagelater{algorithm2e}{2009/11/17}{
1040 | \renewcommand{\algorithmcfname}{\sffamily\bfseries{}Algorithm}%
1041 | \renewcommand{\@algocf@procname}{\sffamily\bfseries{}Procedure}%
1042 | \SetAlgoCaptionSeparator{~}
1043 | \SetAlCapHSkip{0pt}
1044 | \renewcommand{\algocf@captiontext}[2]{%
1045 | \kern0.05em{\color{lipicsYellow}\rule{0.73em}{0.73em}}%
1046 | \hspace*{0.67em}\small #1\algocf@capseparator\nobreakspace#2}
1047 | \renewcommand{\algocf@makecaption}[2]{%
1048 | \parbox[t]{\textwidth}{\algocf@captiontext{#1}{#2}}%
1049 | }%
1050 | \renewcommand{\algocf@captionproctext}[2]{%
1051 | {%
1052 | \kern0.05em{\color{lipicsYellow}\rule{0.73em}{0.73em}}%
1053 | \hspace*{0.67em}\small%
1054 | \ProcSty{\ProcFnt\algocf@procname\ifthenelse{\boolean{algocf@procnumbered}}{\nobreakspace\thealgocf\algocf@typo\algocf@capseparator}{\relax}}%
1055 | \nobreakspace\ProcNameSty{\ProcNameFnt\algocf@captname #2@}% Name of the procedure in ProcName Style.
1056 | \ifthenelse{\equal{\algocf@captparam #2@}{\arg@e}}{}{% if no argument, write nothing
1057 | \ProcNameSty{\ProcNameFnt(}\ProcArgSty{\ProcArgFnt\algocf@captparam #2@}\ProcNameSty{\ProcNameFnt)}%else put arguments in ProcArgSty:
1058 | }% endif
1059 | \algocf@captother #2@%
1060 | }%
1061 | }%
1062 | \renewcommand{\@algocf@capt@boxed}{above}
1063 | \renewcommand{\@algocf@capt@ruled}{above}
1064 | \setlength\algotitleheightrule{0pt}
1065 | }{\ClassWarning{%
1066 | Package algorithm2e outdated}{You are using an outdated version of the package algorithm2e. Please update!}}
1067 | }{}
1068 | \@ifpackageloaded{algorithm}{
1069 | \captionsetup[algorithm]{name=Algorithm, labelformat=boxed, position=top}
1070 | \newcommand\fs@ruled@notop{\def\@fs@cfont{\bfseries}\let\@fs@capt\floatc@ruled
1071 | \def\@fs@pre{}%
1072 | \def\@fs@post{\kern2pt\hrule\relax}%
1073 | \def\@fs@mid{\kern2pt\hrule\kern2pt}%
1074 | \let\@fs@iftopcapt\iftrue}
1075 | \@ifundefined{fst@algorithm}{}{
1076 | \renewcommand\fst@algorithm{\fs@ruled@notop}
1077 | }
1078 | }{}
1079 | \ifx\usecleveref\relax\else
1080 | \@ifpackageloaded{cleveref}{\ClassWarning{Use document option 'cleveref' instead}{Use document option 'cleveref' instead directly loading package 'cleveref'}}{}
1081 | \fi
1082 | \ifx\usethmrestate\relax\else
1083 | \@ifpackageloaded{thm-restate}{\ClassWarning{Use document option 'thm-restate' instead}{Use document option 'thm-restate' instead directly loading package 'thm-restate'}}{}
1084 | \fi
1085 | \ifx\useautoref\relax
1086 | \@ifundefined{algorithmautorefname}{\newcommand{\algorithmautorefname}{Algorithm}}{\renewcommand{\algorithmautorefname}{Algorithm}}%
1087 | \fi
1088 | }
1089 |
1090 | \ifx\usecleveref\relax
1091 | \RequirePackage[capitalise, noabbrev]{cleveref}
1092 | \crefname{algocf}{Algorithm}{Algorithms}
1093 | \Crefname{algocf}{Algorithm}{Algorithms}
1094 | \newcommand{\crefrangeconjunction}{--}
1095 | \newcommand{\creflastconjunction}{, and\nobreakspace}
1096 | \fi
1097 | \ifx\useautoref\relax
1098 | \RequirePackage{aliascnt}
1099 | \fi
1100 | \newtheoremstyle{claimstyle}{\topsep}{\topsep}{}{0pt}{\sffamily}{. }{5pt plus 1pt minus 1pt}%
1101 | {$\vartriangleright$ \thmname{#1}\thmnumber{ #2}\thmnote{ (#3)}}
1102 | \theoremstyle{plain}
1103 | \newtheorem{theorem}{Theorem}
1104 | \ifx\numberwithinsect\relax
1105 | \@addtoreset{theorem}{section}
1106 | \expandafter\def\expandafter\thetheorem\expandafter{%
1107 | \expandafter\thesection\expandafter\@thmcountersep\thetheorem}
1108 | \fi
1109 |
1110 | \ifx\useautoref\relax
1111 | \addto\extrasenglish{%
1112 | \def\chapterautorefname{Chapter}%
1113 | \def\sectionautorefname{Section}%
1114 | \def\subsectionautorefname{Subsection}%
1115 | \def\subsubsectionautorefname{Subsubsection}%
1116 | \def\paragraphautorefname{Paragraph}%
1117 | \def\subparagraphautorefname{Subparagraph}%
1118 | }
1119 | \addto\extrasUKenglish{%
1120 | \def\chapterautorefname{Chapter}%
1121 | \def\sectionautorefname{Section}%
1122 | \def\subsectionautorefname{Subsection}%
1123 | \def\subsubsectionautorefname{Subsubsection}%
1124 | \def\paragraphautorefname{Paragraph}%
1125 | \def\subparagraphautorefname{Subparagraph}%
1126 | }
1127 | \addto\extrasUSenglish{%
1128 | \def\chapterautorefname{Chapter}%
1129 | \def\sectionautorefname{Section}%
1130 | \def\subsectionautorefname{Subsection}%
1131 | \def\subsubsectionautorefname{Subsubsection}%
1132 | \def\paragraphautorefname{Paragraph}%
1133 | \def\subparagraphautorefname{Subparagraph}%
1134 | }
1135 | \ifx\usethmrestate\relax
1136 | \newtheorem{lemma}[theorem]{Lemma}
1137 | \newtheorem{corollary}[theorem]{Corollary}
1138 | \newtheorem{proposition}[theorem]{Proposition}
1139 | \newtheorem{exercise}[theorem]{Exercise}
1140 | \newtheorem{definition}[theorem]{Definition}
1141 | \newtheorem{conjecture}[theorem]{Conjecture}
1142 | \newtheorem{observation}[theorem]{Observation}
1143 | \theoremstyle{definition}
1144 | \newtheorem{example}[theorem]{Example}
1145 | \theoremstyle{remark}
1146 | \newtheorem{note}[theorem]{Note}
1147 | \newtheorem*{note*}{Note}
1148 | \newtheorem{remark}[theorem]{Remark}
1149 | \newtheorem*{remark*}{Remark}
1150 | \theoremstyle{claimstyle}
1151 | \newtheorem{claim}[theorem]{Claim}
1152 | \newtheorem*{claim*}{Claim}
1153 | \else
1154 | \newaliascnt{lemma}{theorem}
1155 | \newtheorem{lemma}[lemma]{Lemma}
1156 | \aliascntresetthe{lemma}
1157 | \newcommand{\lemmaautorefname}{Lemma}
1158 | \newaliascnt{corollary}{theorem}
1159 | \newtheorem{corollary}[corollary]{Corollary}
1160 | \aliascntresetthe{corollary}
1161 | \newcommand{\corollaryautorefname}{Corollary}
1162 | \newaliascnt{proposition}{theorem}
1163 | \newtheorem{proposition}[proposition]{Proposition}
1164 | \aliascntresetthe{proposition}
1165 | \newcommand{\propositionautorefname}{Proposition}
1166 | \newaliascnt{exercise}{theorem}
1167 | \newtheorem{exercise}[exercise]{Exercise}
1168 | \aliascntresetthe{exercise}
1169 | \newcommand{\exerciseautorefname}{Exercise}
1170 | \newaliascnt{definition}{theorem}
1171 | \newtheorem{definition}[definition]{Definition}
1172 | \aliascntresetthe{definition}
1173 | \newcommand{\definitionautorefname}{Definition}
1174 | \newaliascnt{conjecture}{theorem}
1175 | \newtheorem{conjecture}[conjecture]{Conjecture}
1176 | \aliascntresetthe{conjecture}
1177 | \newcommand{\conjectureautorefname}{Conjecture}
1178 | \newaliascnt{observation}{theorem}
1179 | \newtheorem{observation}[observation]{Observation}
1180 | \aliascntresetthe{observation}
1181 | \newcommand{\observationautorefname}{Observation}
1182 | \theoremstyle{definition}
1183 | \newaliascnt{example}{theorem}
1184 | \newtheorem{example}[example]{Example}
1185 | \aliascntresetthe{example}
1186 | \newcommand{\exampleautorefname}{Example}
1187 | \theoremstyle{remark}
1188 | \newaliascnt{note}{theorem}
1189 | \newtheorem{note}[note]{Note}
1190 | \aliascntresetthe{note}
1191 | \newcommand{\noteautorefname}{Note}
1192 | \newtheorem*{note*}{Note}
1193 | \newaliascnt{remark}{theorem}
1194 | \newtheorem{remark}[remark]{Remark}
1195 | \aliascntresetthe{remark}
1196 | \newcommand{\remarkautorefname}{Remark}
1197 | \newtheorem*{remark*}{Remark}
1198 | \theoremstyle{claimstyle}
1199 | \newaliascnt{claim}{theorem}
1200 | \newtheorem{claim}[claim]{Claim}
1201 | \aliascntresetthe{claim}
1202 | \newcommand{\claimautorefname}{Claim}
1203 | \newtheorem*{claim*}{Claim}
1204 | \fi
1205 | \else
1206 | \newtheorem{lemma}[theorem]{Lemma}
1207 | \newtheorem{corollary}[theorem]{Corollary}
1208 | \newtheorem{proposition}[theorem]{Proposition}
1209 | \newtheorem{exercise}[theorem]{Exercise}
1210 | \newtheorem{definition}[theorem]{Definition}
1211 | \newtheorem{conjecture}[theorem]{Conjecture}
1212 | \newtheorem{observation}[theorem]{Observation}
1213 | \theoremstyle{definition}
1214 | \newtheorem{example}[theorem]{Example}
1215 | \theoremstyle{remark}
1216 | \newtheorem{note}[theorem]{Note}
1217 | \newtheorem*{note*}{Note}
1218 | \newtheorem{remark}[theorem]{Remark}
1219 | \newtheorem*{remark*}{Remark}
1220 | \theoremstyle{claimstyle}
1221 | \newtheorem{claim}[theorem]{Claim}
1222 | \newtheorem*{claim*}{Claim}
1223 | \fi
1224 | \theoremstyle{plain}
1225 | \endinput
1226 | %%
1227 | %% End of file `lipics-v2021.cls'.
1228 |
--------------------------------------------------------------------------------