├── .gitignore ├── README.md ├── code ├── Bootstrap.v ├── General.v ├── Makefile ├── Nat.v ├── Prelude.v ├── README.md ├── TestBootstrap │ ├── Nat.v │ └── TiedKnot.v ├── TestGeneral.v └── _CoqProject └── paper ├── cc-by.pdf ├── lipics-logo-bw.pdf ├── lipics-v2019.cls ├── orcid.pdf ├── whynotw.bib └── whynotw.tex /.gitignore: -------------------------------------------------------------------------------- 1 | paper/*.aux 2 | paper/*.bbl 3 | paper/*.blg 4 | paper/*.log 5 | paper/*.out 6 | paper/*.synctex.gz 7 | paper/*.vtc 8 | paper/whynotw.pdf 9 | 10 | code/Makefile.coq 11 | code/Makefile.coq.conf 12 | code/.Makefile.coq.d 13 | code/**/*.vo 14 | code/**/*.vok 15 | code/**/*.vos 16 | code/**/*.glob 17 | code/**/.*.aux 18 | -------------------------------------------------------------------------------- /README.md: -------------------------------------------------------------------------------- 1 | Contrary to previous belief, intensional type theory with W types as the only primitive inductive type is expressive enough to construct the natural numbers along with a whole host of inductive types. 2 | 3 | Coq (version 8.12) sources are in `code`, LaTeX sources are in `paper`. 4 | -------------------------------------------------------------------------------- /code/Bootstrap.v: -------------------------------------------------------------------------------- 1 | (* Author: Jasper Hugunin *) 2 | 3 | From WhyNotW Require Import Prelude. 4 | 5 | (* Everything defined in terms of the primitive type formers only. *) 6 | 7 | #[universes(monomorphic)] Universe UTop. (* For the elimination predicate. *) 8 | 9 | Section universe_of_datatypes. 10 | Universe i. 11 | Universe i'. (* morally i+1 *) 12 | Constraint i < i'. 13 | 14 | Section define_Code. 15 | (* We first have to manually perform the computation of El Code_code *) 16 | Local Notation "#sum_match s 'as' s' 'return' P 'with' f | g 'end'" := 17 | (match s.(fst) as s_fst 18 | return ∀ s_snd, let s' : _ + _ := (s_fst; s_snd) in P 19 | with false => f | true => g end s.(snd)). 20 | Local Notation "#sum_rect [ s 'return' P ]" := 21 | (λ (f : ∀ a, let s : _ + _ := inl a in P) 22 | (g : ∀ b, let s : _ + _ := inr b in P) ↦ 23 | λ x ↦ #sum_match x as s return P with f | g end). 24 | Local Notation sum_rect := (#sum_rect[_ return _]). 25 | Definition Ux1 := prod@{i'} Type@{i} λ _ ↦ 1. (* Fix universe level *) 26 | Definition Code_shapes@{|} : Type@{i'} := (1 ⊔ Ux1 ⊔ 1) ⊔ Ux1 ⊔ 1. 27 | Definition Code_positions@{|} : Code_shapes → Type@{i'} := 28 | #sum_rect [s return Type@{i'}] 29 | (#sum_rect [s return Type@{i'}] 30 | (λ _ ↦ 0) (* nil *) 31 | (#sum_rect [s return Type@{i'}] 32 | (λ (args : Ux1) ↦ sum'@{i'} args.(fst) 0) (* nonind *) 33 | (λ _ ↦ sum'@{i'} 1 (sum'@{i'} 1 0)))) (* choice *) 34 | (#sum_rect [s return Type@{i'}] 35 | (λ _ ↦ sum'@{i'} 1 0) (* inf_ind *) 36 | (λ _ ↦ sum'@{i'} 1 0)). 37 | Definition preCode@{|} : Type@{i'} := W Code_shapes Code_positions. 38 | Definition Code_All@{j | i <= j} {s : Code_shapes} 39 | (C : Code_positions s → Type@{j}) : Type@{j} := 40 | let positions := Code_positions in 41 | (#sum_rect [a return (positions a → Type@{j}) → Type@{j}] 42 | (#sum_rect [a return (positions (inl a) → Type@{j}) → Type@{j}] 43 | (λ _ _ ↦ 1) 44 | (#sum_rect [a return (positions (inl (inr a)) → Type@{j}) → Type@{j}] 45 | (λ args C ↦ (∀ i, C (inl i)) × 1) 46 | (λ args C ↦ C None × C (Some None) × 1))) 47 | (#sum_rect [a return (positions (inr a) → Type@{j}) → Type@{j}] 48 | (λ args C ↦ C None × 1) 49 | (λ args C ↦ C None × 1)) 50 | s C). 51 | Definition Code_F@{j | i < j} (X : Type@{j}) : Type@{j} := 52 | (1 ⊔ (* nil *) 53 | (prod@{j} Type@{i} λ A ↦ (A → X) × 1) ⊔ (* nonind *) 54 | (X × X × 1)) ⊔ (* choice *) 55 | (Type@{i} × X × 1) ⊔ (* inf_ind *) 56 | (X × 1). (* ind *) 57 | Definition Code_Fmap@{j1 j2} {X : Type@{j1}} {Y : Type@{j2}} (f : X → Y) : 58 | Code_F@{j1} X → Code_F@{j2} Y := 59 | #sum_rect [s return Code_F@{j2} Y] 60 | (inl' ∘ #sum_rect [s return (1 ⊔ (prod@{j2} Type@{i} λ A ↦ (A → Y) × 1) ⊔ (Y × Y × 1))] 61 | (λ x ↦ inl' x) (* nil *) 62 | (inr' ∘ #sum_rect [s return (prod@{j2} Type@{i} λ A ↦ (A → Y) × 1) ⊔ (Y × Y × 1)] 63 | (λ args : prod@{j1} Type@{i} λ A ↦ (A → X) × 1 ↦ (* nonind *) 64 | inl' (A := prod@{j2} Type@{i} λ A ↦ (A → Y) × 1) 65 | (args.(fst); f ∘ args.(snd).(fst); args.(snd).(snd))) 66 | (λ args : X × X × 1 ↦ (* choice *) 67 | inr' (f args.(fst); f args.(snd).(fst); args.(snd).(snd))))) 68 | (inr' ∘ #sum_rect [s return (Type@{i} × Y × 1) ⊔ (Y × 1)] 69 | (λ args : Type@{i} × X × 1 ↦ (* inf_ind *) 70 | inl'@{j2} (A := Type@{i} × Y × 1) 71 | (args.(fst); f args.(snd).(fst); args.(snd).(snd))) 72 | (λ args : X × 1 ↦ inr'@{j2} (f args.(fst); args.(snd)))). (* ind *) 73 | Definition Code_get_shape@{j} {X : Type@{j}} : Code_F@{j} X → Code_shapes := 74 | #sum_rect [_ return Code_shapes] 75 | (λ t ↦ inl'@{i'} (#sum_rect [_ return 1 ⊔ Ux1 ⊔ 1] 76 | (λ args ↦ inl'@{i'} args) (* nil *) 77 | (λ t ↦ inr'@{i'} (#sum_rect [_ return Ux1 ⊔ 1] 78 | (λ (args : prod@{j} Type@{i} λ _ ↦ _ × 1) ↦ 79 | inl'@{i'} (A := Ux1) (args.(fst); args.(snd).(snd))) (* nonind *) 80 | (λ (args : X × X × 1) ↦ inr'@{i'} args.(snd).(snd)) (* choice *) 81 | t)) 82 | t)) 83 | (λ t ↦ inr@{i'} (#sum_rect [_ return Ux1 ⊔ 1] 84 | (λ args : prod@{j} Type@{i} λ _ ↦ prod@{j} _ λ _ ↦ 1 ↦ 85 | inl@{i'} (args.(fst); args.(snd).(snd))) (* inf_ind *) 86 | (λ args ↦ inr'@{i'} args.(snd)) (* ind *) 87 | t)). 88 | Definition Code_get_f@{j} {X : Type@{j}} : 89 | ∀ fx : Code_F@{j} X, Code_positions (Code_get_shape fx) → X := 90 | let get_shape := Code_get_shape in 91 | let positions := Code_positions in 92 | #sum_rect [t return positions (get_shape t) → X] 93 | (#sum_rect [t return positions (get_shape (inl t)) → X] 94 | (λ _ (XX : 0) ↦ match XX with end) (* nil *) 95 | (#sum_rect [t return positions (get_shape (inl (inr t))) → X] 96 | (λ args ↦ sum_rect (* nonind *) 97 | (args.(snd).(fst)) 98 | (λ XX : 0 ↦ match XX with end)) 99 | (λ args ↦ sum_rect (* choice *) 100 | (λ '★ ↦ args.(fst)) 101 | (sum_rect 102 | (λ '★ ↦ args.(snd).(fst)) 103 | (λ XX : 0 ↦ match XX with end))))) 104 | (#sum_rect [t return positions (get_shape (inr t)) → X] 105 | (λ args ↦ sum_rect (* inf_ind *) 106 | (λ '★ ↦ args.(snd).(fst)) 107 | (λ XX : 0 ↦ match XX with end)) 108 | (λ args ↦ sum_rect (* ind *) 109 | (λ '★ ↦ args.(fst)) 110 | (λ XX : 0 ↦ match XX with end))). 111 | Fixpoint Code_canonical@{} (x : preCode) : Type@{i'} := match x with 112 | | sup a f => prod@{i'} 113 | (Code_All (Code_canonical ∘ f)) λ _ ↦ 114 | prod@{i'} 115 | (Code_F preCode) λ t ↦ 116 | (Id (A := Σ (a : Code_shapes), Code_positions a → preCode) 117 | (Code_get_shape t; Code_get_f t) (a; f)) 118 | end. 119 | 120 | Section split. 121 | Context 122 | {X : Type@{i'}} 123 | {C : X → Type@{i'}} 124 | . 125 | Let T := prod@{i'} X C. 126 | Definition Code_get_heredity 127 | (Hered := λ (x : Code_F T) ↦ 128 | let x_fst := Code_Fmap fst x in 129 | Code_All (s := Code_get_shape x_fst) (C ∘ Code_get_f x_fst) : Type@{i'}) : 130 | ∀ x, Hered x := 131 | #sum_rect [x return Hered x] 132 | (#sum_rect [x return Hered (inl x)] 133 | (λ x ↦ x) 134 | (#sum_rect [x return Hered (inl (inr x))] 135 | (λ x : prod@{i'} Type@{i} λ A ↦ (A → T) × 1 ↦ 136 | (snd ∘ x.(snd).(fst); x.(snd).(snd))) 137 | (λ x : T × T × 1 ↦ (snd x.(fst); snd x.(snd).(fst); x.(snd).(snd))))) 138 | (#sum_rect [x return Hered (inr x)] 139 | (λ x : prod@{i'} Type@{i} λ Ix ↦ T × 1 ↦ 140 | (snd x.(snd).(fst); x.(snd).(snd))) 141 | (λ x : T × 1 ↦ (snd x.(fst); x.(snd)))). 142 | 143 | Definition Fiber {A B1 : Type@{i'}} (B2 : B1 → Type@{i'}) 144 | (f : A → B1) (g : ∀ a, B2 (f a)) (b1 : B1) (b2 : B2 b1) : Type@{i'} := 145 | Σ (a : A), Id (A := Σ (b1 : B1), B2 b1) (f a; g a) (b1; b2). 146 | Definition Code_split_view 147 | (x : Code_F X) (hered : Code_All (C ∘ Code_get_f x)) : Type@{i'} := 148 | Fiber (λ x ↦ Code_All (C ∘ Code_get_f x)) 149 | (Code_Fmap fst) (Code_get_heredity) x hered. 150 | 151 | Section choice. 152 | Context 153 | {Al Bl Sl : Type@{i'}} {Pl : Sl → Type@{i'}} 154 | {Alll : ∀ s, (Pl s → Type@{i'}) → Type@{i'}} 155 | {fl : Al → Bl} {gsl : Bl → Sl} {gfl : ∀ b, Pl (gsl b) → X} 156 | {gl : ∀ a, Alll (gsl (fl a)) (C ∘ gfl (fl a))} 157 | {Ar Br Sr : Type@{i'}} {Pr : Sr → Type@{i'}} 158 | {Allr : ∀ s, (Pr s → Type@{i'}) → Type@{i'}} 159 | {fr : Ar → Br} {gsr : Br → Sr} {gfr : ∀ b, Pr (gsr b) → X} 160 | {gr : ∀ a, Allr (gsr (fr a)) (C ∘ gfr (fr a))} 161 | . 162 | Let A : Type@{i'} := Al ⊔ Ar. 163 | Let B : Type@{i'} := Bl ⊔ Br. 164 | Let S : Type@{i'} := Sl ⊔ Sr. 165 | Let P : S → Type@{i'} := sum_rect Pl Pr. 166 | Let All : ∀ s, (P s → Type@{i'}) → Type@{i'} := 167 | #sum_rect [s return (P s → Type@{i'}) → Type@{i'}] Alll Allr. 168 | Let f : A → B := sum_rect (inl' ∘ fl) (inr' ∘ fr). 169 | Let gs : B → S := sum_rect (inl' ∘ gsl) (inr' ∘ gsr). 170 | Let gf : ∀ b, P (gs b) → X := #sum_rect [b return P (gs b) → X] gfl gfr. 171 | Let g : ∀ a, All (gs (f a)) (C ∘ gf (f a)) := 172 | #sum_rect [a return All (gs (f a)) (C ∘ gf (f a))] gl gr. 173 | Definition Code_split_view_cover_choice_l a hered : 174 | Fiber (λ x ↦ Alll _ (C ∘ gfl x)) fl gl a hered → 175 | Fiber (λ x ↦ All _ (C ∘ gf x)) f g (inl a) hered := 176 | λ cover ↦ match cover with (xc; e) => 177 | match e in Id _ (a; hered) return Fiber _ f g (inl a) hered 178 | with refl => (inl xc; refl) end 179 | end. 180 | Definition Code_split_view_cover_choice_r b hered : 181 | Fiber (λ x ↦ Allr _ (C ∘ gfr x)) fr gr b hered → 182 | Fiber (λ x ↦ All _ (C ∘ gf x)) f g (inr b) hered := 183 | λ cover ↦ match cover with (xc; e) => 184 | match e in Id _ (b; hered) return Fiber _ f g (inr b) hered 185 | with refl => (inr xc; refl) end 186 | end. 187 | Definition Code_split_view_cover_choice@{|} : 188 | (∀ a hered, Fiber (λ x ↦ Alll _ (C ∘ gfl x)) fl gl a hered) → 189 | (∀ b hered, Fiber (λ x ↦ Allr _ (C ∘ gfr x)) fr gr b hered) → 190 | ∀ x hered, Fiber (λ x ↦ All _ (C ∘ gf x)) f g x hered := 191 | λ Hl Hr ↦ #sum_rect [x return ∀ hered, Fiber (λ x ↦ All _ (C ∘ gf x)) f g x hered] 192 | (λ a hered ↦ Code_split_view_cover_choice_l a hered (Hl a hered)) 193 | (λ b hered ↦ Code_split_view_cover_choice_r b hered (Hr b hered)). 194 | End choice. 195 | Arguments Code_split_view_cover_choice 196 | {Al Bl Sl Pl Alll fl gsl gfl gl} {Ar Br Sr Pr Allr fr gsr gfr gr} & Hl Hr x hered. 197 | Definition Code_split_view_cover : ∀ x hered, Code_split_view x hered := 198 | Code_split_view_cover_choice 199 | (Code_split_view_cover_choice 200 | (λ '[] '[] ↦ ([]; refl)) 201 | (Code_split_view_cover_choice 202 | (λ '[A; B] '[BH] ↦ ([A; λ a ↦ (B a; BH a)]; refl)) 203 | (λ '[A; B] '[AH; BH] ↦ ([(A; AH); (B; BH)]; refl)))) 204 | (Code_split_view_cover_choice 205 | (λ '[Ix; B] '[BH] ↦ ([Ix; (B; BH)]; refl)) 206 | (λ '[B] '[BH] ↦ ([(B; BH)]; refl))). 207 | End split. 208 | 209 | (* This is the true definition of Code. *) 210 | Definition Code@{|} : Type@{i'} := Σ (x : preCode), Code_canonical x. 211 | Definition Code_intro@{|} (x : Code_F Code) : Code := 212 | let fst_x := Code_Fmap fst x in 213 | (sup _ _; Code_get_heredity x; fst_x; refl). 214 | 215 | Local Notation nil' x := (inl (inl x)). 216 | Local Notation nonind' x := (inl (inr (inl x))). 217 | Local Notation choice' x := (inl (inr (inr x))). 218 | Local Notation inf_ind' x := (inr (inl x)). 219 | Local Notation ind' x := (inr (inr x)). 220 | 221 | Definition nil : Code := Code_intro (nil' []). 222 | Definition nonind (A : Type@{i}) (B : A → Code) : Code := 223 | Code_intro (nonind' [A; B]). 224 | Definition inf_ind (Ix : Type@{i}) (B : Code) : Code := 225 | Code_intro (inf_ind' [Ix; B]). 226 | Definition ind (B : Code) : Code := 227 | Code_intro (ind' [B]). 228 | Definition choice (A B : Code) : Code := 229 | Code_intro (choice' [A; B]). 230 | 231 | (* Recursion for Code *) 232 | Section Code_rec. 233 | Constraint i <= UTop. 234 | Context 235 | (P : Code → Type@{UTop}) 236 | . 237 | Section general_rec. 238 | Context 239 | (IS : ∀ (x : Code_F Code), Code_All@{UTop} (P ∘ Code_get_f x) → P (Code_intro x)) 240 | . 241 | 242 | Section build_IH. 243 | Let Code_All_j {s} := Code_All@{UTop} (s := s). 244 | Local Notation R x := 245 | (∀ (IH : ∀ p c, P (Code_get_f (Code_Fmap fst x) p; c)), 246 | Code_All_j (P ∘ Code_get_f x)). 247 | Definition Code_build_IH@{|} : 248 | ∀ x, R x := 249 | #sum_rect [x return R x] 250 | (#sum_rect [x return R (inl x)] 251 | (λ x IH ↦ x) (* nil *) 252 | (#sum_rect [x return R (inl (inr x))] 253 | (λ x IH ↦ (* nonind *) 254 | (λ i ↦ IH (inl i) (x.(snd).(fst) i).(snd); x.(snd).(snd))) 255 | (λ x IH ↦ (* choice *) 256 | (IH None x.(fst).(snd); IH (Some None) x.(snd).(fst).(snd); 257 | x.(snd).(snd))))) 258 | (#sum_rect [x return R (inr x)] 259 | (λ x IH ↦ (* inf_ind *) 260 | (IH None x.(snd).(fst).(snd); x.(snd).(snd))) 261 | (λ x IH ↦ (* ind *) 262 | (IH None x.(fst).(snd); x.(snd)))). 263 | End build_IH. 264 | 265 | Definition Code_rec_helper@{|} t (hered : Code_All (Code_canonical ∘ _)) 266 | (IH : ∀ p c, P (Code_get_f t p; c)) : 267 | Code_split_view t hered → P (sup _ _; hered; t; refl) := 268 | fun cover => match cover with (x; e) => 269 | match e in Id _ (t; hered) 270 | return (∀ p c, P (Code_get_f t p; c)) → P (sup _ _; hered; t; refl) 271 | with refl => λ IH ↦ 272 | IS x (Code_build_IH x IH) 273 | end 274 | end IH. 275 | Definition Code_rec_gen@{|} : ∀ A, P A := 276 | let fix rec x : ∀ c, P (x ; c) := match x with 277 | | sup s f => λ '((hered; t; eq) : Code_canonical (sup s f)) ↦ 278 | match eq in Id _ (s; f) 279 | return 280 | ∀ hered : Code_All@{i'} (Code_canonical ∘ f), 281 | (∀ p c, P (f p; c)) → 282 | P (sup s f; hered; t; eq) 283 | with refl => λ hered IH ↦ 284 | Code_rec_helper t hered IH (Code_split_view_cover t hered) 285 | end hered (rec ∘ f) 286 | end in 287 | λ A ↦ rec A.(fst) A.(snd). 288 | End general_rec. 289 | 290 | Context 291 | (IS_nil : P nil) 292 | (IS_nonind : ∀ (A : Type@{i}) B, (∀ a, P (B a)) → P (nonind A B)) 293 | (IS_inf_ind : ∀ (Ix : Type@{i}) B, P B → P (inf_ind Ix B)) 294 | (IS_ind : ∀ B, P B → P (ind B)) 295 | (IS_choice : ∀ A B, P A → P B → P (choice A B)) 296 | . 297 | 298 | Definition IS@{|} (arg : Code_F@{i'} Code) : 299 | Code_All@{UTop} (P ∘ Code_get_f arg) → P (Code_intro arg) := 300 | match arg return Code_All@{UTop} (P ∘ Code_get_f arg) → P (Code_intro arg) with 301 | | nil' [] => λ '[] ↦ IS_nil 302 | | nonind' [A; B] => λ '[IH] ↦ IS_nonind A B IH 303 | | inf_ind' [A; B] => λ '[IH] ↦ IS_inf_ind A B IH 304 | | ind' [B] => λ '[IH] ↦ IS_ind B IH 305 | | choice' [A; B] => λ '[IHA; IHB] ↦ IS_choice A B IHA IHB 306 | end. 307 | 308 | Definition Code_rec@{|} := Code_rec_gen IS. 309 | End Code_rec. 310 | Arguments Code_rec P & IS_nil IS_nonind IS_inf_ind IS_ind IS_choice. 311 | End define_Code. 312 | 313 | (* ---------------------------------------------------------------- *) 314 | (* 315 | We have now completed the definition of Code from W. 316 | The rest follows General.v, but is *very* slow to typecheck. 317 | *) 318 | 319 | Definition F@{j | i <= j, j < UTop} : Code → Type@{j} → Type@{j} := 320 | Code_rec (λ A ↦ Type@{j} → Type@{j}) 321 | (λ X ↦ 1) 322 | (λ A B F_B ↦ λ X ↦ Σ (a : A), (F_B a) X) 323 | (λ Ix B F_B ↦ λ X ↦ (Ix → X) × F_B X) 324 | (λ B F_B ↦ λ X ↦ X × F_B X) 325 | (λ A B F_A F_B ↦ λ X ↦ F_A X ⊔ F_B X). 326 | 327 | Definition Fmap@{j1 j2} 328 | (A : Code) {X : Type@{j1}} {Y : Type@{j2}} (f : X → Y) 329 | : F@{j1} A X → F@{j2} A Y := 330 | Code_rec (λ A ↦ F@{j1} A X → F@{j2} A Y) 331 | (λ x ↦ x) 332 | (λ A B Fmap_B ↦ λ x ↦ (x.(fst); (Fmap_B x.(fst)) x.(snd))) 333 | (λ Ix B Fmap_B ↦ λ x ↦ (f ∘ x.(fst); Fmap_B x.(snd))) 334 | (λ B Fmap_B ↦ λ x ↦ (f x.(fst); Fmap_B x.(snd))) 335 | (λ A B Fmap_A Fmap_B ↦ λ x ↦ match x with 336 | | inl a => inl (Fmap_A a) 337 | | inr b => inr (Fmap_B b) 338 | end) 339 | A. 340 | 341 | (* Checking lawfulness is left to the reader *) 342 | 343 | (* We next define the standard construction in W types. *) 344 | 345 | Definition shapes : Code → Type@{i} := 346 | Code_rec (λ A ↦ Type@{i}) 347 | 1 348 | (λ A B shapes_B ↦ Σ (a : A), shapes_B a) 349 | (λ Ix B shapes_B ↦ shapes_B) 350 | (λ B shapes_B ↦ shapes_B) 351 | (λ A B shapes_A shapes_B ↦ shapes_A ⊔ shapes_B). 352 | 353 | Definition positions : ∀ A, shapes A → Type@{i} := 354 | Code_rec (λ A ↦ shapes A → Type@{i}) 355 | (λ x ↦ 0) 356 | (λ A B positions_B ↦ λ x ↦ positions_B x.(fst) x.(snd)) 357 | (λ Ix B positions_B ↦ λ x ↦ Ix ⊔ positions_B x) 358 | (λ B positions_B ↦ λ x ↦ 1 ⊔ positions_B x) 359 | (λ A B positions_A positions_B ↦ λ x ↦ match x with 360 | | inl a => positions_A a 361 | | inr b => positions_B b 362 | end). 363 | 364 | Definition get_shape@{j} {A : Code} {X : Type@{j}} : F@{j} A X → shapes A := 365 | Code_rec (λ A ↦ F@{j} A X → shapes A) 366 | (λ x ↦ x) 367 | (λ A B get_shape_B ↦ λ x ↦ (x.(fst); get_shape_B _ x.(snd))) 368 | (λ Ix B get_shape_B ↦ λ x ↦ get_shape_B x.(snd)) 369 | (λ B get_shape_B ↦ λ x ↦ get_shape_B x.(snd)) 370 | (λ A B get_shape_A get_shape_B ↦ λ x ↦ match x with 371 | | inl a => inl (get_shape_A a) 372 | | inr b => inr (get_shape_B b) 373 | end) 374 | A. 375 | 376 | Definition get_f@{j} {A : Code} {X : Type@{j}} : 377 | ∀ x : F@{j} A X, positions A (get_shape x) → X := 378 | Code_rec (λ A ↦ ∀ x : F@{j} A X, positions A (get_shape x) → X) 379 | (λ _ XX ↦ match XX with end) 380 | (λ A B get_f_B ↦ λ x ↦ get_f_B _ x.(snd)) 381 | (λ Ix B get_f_B ↦ λ x p ↦ match p with 382 | | inl i => x.(fst) i 383 | | inr p => get_f_B x.(snd) p 384 | end) 385 | (λ B get_f_B ↦ λ x p ↦ match p with 386 | | None => x.(fst) 387 | | Some p => get_f_B x.(snd) p 388 | end) 389 | (λ A B get_f_A get_f_B ↦ λ x ↦ match x with 390 | | inl a => get_f_A a 391 | | inr b => get_f_B b 392 | end) 393 | A. 394 | 395 | Definition get_W_arg@{j} {A : Code} {X : Type@{j}} (x : F@{j} A X) : 396 | Σ (s : shapes A), (positions A s → X) := 397 | (get_shape x; get_f x). 398 | 399 | (* These functors induce a refined notion of forall p : positions A s, P p *) 400 | Definition All@{j} : ∀ A {s}, (positions A s → Type@{j}) → Type@{j} := 401 | Code_rec (λ A ↦ ∀ s, (positions A s → Type@{j}) → Type@{j}) 402 | (λ _ _ ↦ 1) 403 | (λ A B All_B ↦ λ s P ↦ All_B s.(fst) _ P) 404 | (λ Ix B All_B ↦ λ s P ↦ (∀ i, P (inl i)) × All_B _ (P ∘ inr')) 405 | (λ B All_B ↦ λ s P ↦ P None × All_B _ (P ∘ Some')) 406 | (λ A B All_A All_B ↦ λ s ↦ match s with 407 | | inl a => All_A a 408 | | inr b => All_B b 409 | end). 410 | 411 | Definition All_in@{j} : 412 | ∀ A {s} {P : positions A s → Type@{j}}, 413 | (∀ p, P p) → All@{j} A (s := s) P := 414 | Code_rec 415 | (λ A ↦ ∀ s (P : positions A s → Type@{j}), 416 | (∀ p, P p) → All@{j} A (s := s) P) 417 | (λ s _ _ ↦ s) 418 | (λ A B All_in_B ↦ λ s P f ↦ All_in_B s.(fst) _ _ f) 419 | (λ Ix B All_in_B ↦ λ s P f ↦ (f ∘ inl'; All_in_B _ _ (f ∘ inr'))) 420 | (λ B All_in_B ↦ λ s P f ↦ (f None; All_in_B _ _ (f ∘ Some'))) 421 | (λ A B All_in_A All_in_B ↦ λ s ↦ match s with 422 | | inl a => All_in_A a 423 | | inr b => All_in_B b 424 | end). 425 | 426 | (* We can lift predicates over these functors: *) 427 | Definition liftP@{j k} (A : Code) {X : Type@{j}} (P : X → Type@{k}) 428 | (x : F@{j} A X) : Type@{k} := 429 | All@{k} A (P ∘ get_f x). 430 | 431 | Definition preEl (A : Code) : Type@{i} := W (shapes A) (positions A). 432 | 433 | (* 434 | Then we define the canonical predicate. 435 | A term is canonical if its subterms are canonical 436 | and (s; f) is in the image of get_W_arg. 437 | *) 438 | Definition canonical@{} (A : Code) : preEl A → Type@{i} := 439 | fix canonical x := match x with 440 | | sup s f => 441 | All A (canonical ∘ f) × 442 | Σ (t : F A (preEl A)), Id (get_W_arg t) (s; f) 443 | end. 444 | 445 | Definition El (A : Code) : Type@{i} := Σ (x : preEl A), canonical A x. 446 | 447 | Section split. 448 | (* 449 | We prove an equivalence between F A (Σ (x : X), C x) and 450 | Σ (fx : F A X), All A (C ∘ get_f x) 451 | *) 452 | Context 453 | {X : Type@{i}} 454 | {C : X → Type@{i}} 455 | . 456 | Let T : Type@{i} := Σ (x : X), C x. 457 | 458 | Definition get_heredity : ∀ {A : Code}, 459 | ∀ (x : F A T), 460 | let x_fst := Fmap A fst x in 461 | All A (s := get_shape x_fst) (C ∘ get_f x_fst) := 462 | Code_rec 463 | (λ A ↦ ∀ (x : F A T), 464 | let x_fst := Fmap A fst x in 465 | All A (s := get_shape x_fst) (C ∘ get_f x_fst)) 466 | (λ x ↦ x) 467 | (λ A B get_heredity_B ↦ λ x ↦ get_heredity_B x.(fst) x.(snd)) 468 | (λ Ix B get_heredity_B ↦ λ x ↦ (snd ∘ x.(fst); get_heredity_B x.(snd))) 469 | (λ B get_heredity_B ↦ λ x ↦ (snd x.(fst); get_heredity_B x.(snd))) 470 | (λ A B get_heredity_A get_heredity_B ↦ λ x ↦ match x with 471 | | inl a => get_heredity_A a 472 | | inr b => get_heredity_B b 473 | end). 474 | 475 | (* The pair (Fmap fst, get_heredity) is right-invertible *) 476 | Definition split_view (A : Code) (x : F A X) (hered : All A (C ∘ get_f x)) : 477 | Type@{i} := 478 | Σ (xc : F A T), 479 | Id (A := Σ (x : F A X), All A (C ∘ get_f x)) 480 | (Fmap A fst xc; get_heredity xc) (x; hered). 481 | Definition split_view_refl (A : Code) (xc : F A T) : 482 | split_view A (Fmap A fst xc) (get_heredity xc) := 483 | (xc; refl). 484 | Definition split_view_elim@{j} {A} (P : ∀ x hered, split_view A x hered → Type@{j}) 485 | (H : ∀ xc, P (Fmap A fst xc) (get_heredity xc) (split_view_refl A xc)) x hered 486 | (cover : split_view A x hered) : P x hered cover := 487 | match cover.(snd) as e in Id _ (x; hered) return P x hered (cover.(fst); e) 488 | with refl => H cover.(fst) end. 489 | Definition split_view_cover_nonind A B x hered : 490 | split_view (B x.(fst)) x.(snd) hered → split_view (nonind A B) x hered := 491 | split_view_elim@{i} 492 | (λ x_snd hered _ ↦ split_view (nonind A B) (x.(fst); x_snd) hered) 493 | (λ xc ↦ split_view_refl (nonind A B) (x.(fst); xc)) 494 | x.(snd) hered. 495 | Definition split_view_cover_inf_ind Ix B x hered : 496 | split_view B x.(snd) hered.(snd) → split_view (inf_ind Ix B) x hered := 497 | split_view_elim@{i} 498 | (λ x_snd hered_snd _ ↦ 499 | split_view (inf_ind Ix B) (x.(fst); x_snd) (hered.(fst); hered_snd)) 500 | (λ xc ↦ split_view_refl (inf_ind Ix B) (λ i ↦ (x.(fst) i; hered.(fst) i); xc)) 501 | x.(snd) hered.(snd). 502 | Definition split_view_cover_ind B x hered : 503 | split_view B x.(snd) hered.(snd) → split_view (ind B) x hered := 504 | split_view_elim@{i} 505 | (λ x_snd hered_snd _ ↦ 506 | split_view (ind B) (x.(fst); x_snd) (hered.(fst); hered_snd)) 507 | (λ xc ↦ split_view_refl (ind B) ((x.(fst); hered.(fst)); xc)) 508 | x.(snd) hered.(snd). 509 | Definition split_view_cover_choice_l A B a hered : 510 | split_view A a hered → split_view (choice A B) (inl a) hered := 511 | split_view_elim@{i} 512 | (λ a hered _ ↦ split_view (choice A B) (inl a) hered) 513 | (λ xc ↦ split_view_refl (choice A B) (inl xc)) 514 | a hered. 515 | Definition split_view_cover_choice_r A B b hered : 516 | split_view B b hered → split_view (choice A B) (inr b) hered := 517 | split_view_elim@{i} 518 | (λ b hered _ ↦ split_view (choice A B) (inr b) hered) 519 | (λ xc ↦ split_view_refl (choice A B) (inr xc)) 520 | b hered. 521 | Definition split_view_cover : ∀ (A : Code) x hered, split_view A x hered := 522 | Code_rec (λ A ↦ ∀ x hered, split_view A x hered) 523 | (λ '[] '[] ↦ split_view_refl nil []) 524 | (λ A B split_view_cover_B ↦ λ x hered ↦ 525 | split_view_cover_nonind A B _ _ (split_view_cover_B x.(fst) x.(snd) hered)) 526 | (λ Ix B split_view_cover_B ↦ λ x hered ↦ 527 | split_view_cover_inf_ind Ix B _ _ (split_view_cover_B x.(snd) hered.(snd))) 528 | (λ B split_view_cover_B ↦ λ x hered ↦ 529 | split_view_cover_ind B _ _ (split_view_cover_B x.(snd) hered.(snd))) 530 | (λ A B split_view_cover_A split_view_cover_B ↦ λ x ↦ match x with 531 | | inl a => λ hered ↦ 532 | split_view_cover_choice_l A B _ _ (split_view_cover_A a hered) 533 | | inr b => λ hered ↦ 534 | split_view_cover_choice_r A B _ _ (split_view_cover_B b hered) 535 | end). 536 | 537 | (* And the fibers are in fact contractible *) 538 | Definition split_view_isContr : 539 | ∀ A x hered other, Id other (split_view_cover A x hered) := 540 | Code_rec (λ A ↦ ∀ x hered other, Id other (split_view_cover A x hered)) 541 | (λ x hered ↦ split_view_elim@{i} 542 | (λ x hered other ↦ Id other (split_view_cover nil x hered)) 543 | (λ '[] ↦ refl) 544 | _ _) 545 | (λ A B split_view_isContr_B ↦ λ x hered ↦ split_view_elim@{i} 546 | (λ x hered other ↦ Id other (split_view_cover (nonind A B) x hered)) 547 | (λ xc ↦ cong (split_view_cover_nonind A B (Fmap _ fst xc) (get_heredity xc)) 548 | (split_view_isContr_B _ _ _ (split_view_refl _ xc.(snd)))) 549 | _ _) 550 | (λ Ix B split_view_isContr_B ↦ λ x hered ↦ split_view_elim@{i} 551 | (λ x hered other ↦ Id other (split_view_cover (inf_ind Ix B) x hered)) 552 | (λ xc ↦ cong (split_view_cover_inf_ind Ix B (Fmap _ fst xc) (get_heredity xc)) 553 | (split_view_isContr_B _ _ (split_view_refl _ xc.(snd)))) 554 | _ _) 555 | (λ B split_view_isContr_B ↦ λ x hered ↦ split_view_elim@{i} 556 | (λ x hered other ↦ Id other (split_view_cover (ind B) x hered)) 557 | (λ xc ↦ cong (split_view_cover_ind B (Fmap _ fst xc) (get_heredity xc)) 558 | (split_view_isContr_B _ _ (split_view_refl _ xc.(snd)))) 559 | _ _) 560 | (λ A B split_view_isContr_A split_view_isContr_B ↦ λ x hered ↦ 561 | split_view_elim@{i} 562 | (λ x hered other ↦ Id other (split_view_cover (choice A B) x hered)) 563 | (λ xc ↦ match xc with 564 | | inl a => 565 | cong (split_view_cover_choice_l A B (Fmap _ fst a) (get_heredity a)) 566 | (split_view_isContr_A _ _ (split_view_refl _ a)) 567 | | inr b => 568 | cong (split_view_cover_choice_r A B (Fmap _ fst b) (get_heredity b)) 569 | (split_view_isContr_B _ _ (split_view_refl _ b)) 570 | end) 571 | _ _). 572 | End split. 573 | 574 | Definition intro (A : Code) : F@{i} A (El A) → El A := 575 | λ x ↦ let pre_x := Fmap A fst x in 576 | (sup (get_shape pre_x) (get_f pre_x); get_heredity x; pre_x; refl). 577 | 578 | (* Now for induction *) 579 | 580 | Section induction. 581 | Universe j. 582 | Context 583 | (A : Code) 584 | (P : El A → Type@{j}) 585 | (IS : ∀ (x : F A (El A)), liftP@{i j} A P x → P (intro A x)) 586 | . 587 | 588 | Definition build_IH@{} : 589 | ∀ A' x (IH : ∀ p c, P (get_f (Fmap@{i i} A' fst x) p; c)), 590 | All@{j} A' (P ∘ get_f x) := 591 | Code_rec 592 | (λ A' ↦ ∀ x (IH : ∀ p c, P (get_f (Fmap@{i i} A' fst x) p; c)), 593 | All@{j} A' (P ∘ get_f x)) 594 | (λ x IH ↦ x) 595 | (λ A B build_IH_B ↦ λ x IH ↦ build_IH_B x.(fst) x.(snd) IH) 596 | (λ Ix B build_IH_B ↦ λ x IH ↦ 597 | (λ i ↦ IH (inl i) (x.(fst) i).(snd); build_IH_B x.(snd) (IH ∘ inr'))) 598 | (λ B build_IH_B ↦ λ x IH ↦ 599 | (IH None x.(fst).(snd); build_IH_B x.(snd) (IH ∘ Some'))) 600 | (λ A B build_IH_A build_IH_B ↦ λ x ↦ match x with 601 | | inl a => build_IH_A a 602 | | inr b => build_IH_B b 603 | end). 604 | 605 | Definition build_IH_eq@{} (A' : Code) (el : ∀ x, P x) : 606 | ∀ x, Id (build_IH A' x (λ p c ↦ el _)) (All_in A' (λ p ↦ el _)) := 607 | Code_rec 608 | (λ A' ↦ ∀ x, Id (build_IH A' x (λ p c ↦ el _)) (All_in A' (λ p ↦ el _))) 609 | (λ x ↦ refl) 610 | (λ A B build_IH_eq_B ↦ λ x ↦ build_IH_eq_B x.(fst) x.(snd)) 611 | (λ Ix B build_IH_eq_B ↦ λ x ↦ 612 | cong (pair (el ∘ x.(fst))) (build_IH_eq_B x.(snd))) 613 | (λ B build_IH_eq_B ↦ λ x ↦ 614 | cong (pair (el x.(fst))) (build_IH_eq_B x.(snd))) 615 | (λ A B build_IH_eq_A build_IH_eq_B ↦ λ x ↦ match x with 616 | | inl a => build_IH_eq_A a 617 | | inr b => build_IH_eq_B b 618 | end) 619 | A'. 620 | 621 | Definition rec_helper@{} t (hered : All A (canonical A ∘ _)) 622 | (IH : ∀ p c, P (get_f t p; c)) : 623 | split_view A t hered → P (sup _ _; hered; t; refl) := 624 | λ cover ↦ split_view_elim@{j} 625 | (λ t hered _ ↦ (∀ p c, P (get_f t p; c)) → P (sup _ _; hered; t; refl)) 626 | (λ x ↦ λ IH ↦ IS x (build_IH A x IH)) 627 | _ _ cover IH. 628 | Definition rec@{} : ∀ x, P x := 629 | let fix rec (x : preEl A) : ∀ c, P (x; c) := match x with 630 | | sup s f => λ '((hered; t; eq) : canonical A (sup s f)) ↦ 631 | match eq in Id _ (s; f) 632 | return 633 | ∀ hered : All A (canonical A ∘ f), 634 | (∀ p c, P (f p; c)) → 635 | P (sup s f; hered; t; eq) 636 | with refl => λ hered IH ↦ 637 | rec_helper t hered IH (split_view_cover A t hered) 638 | end hered (λ p ↦ rec (f p)) 639 | end in 640 | λ x ↦ rec x.(fst) x.(snd). 641 | 642 | Definition rec_eq@{} x 643 | : Id (rec (intro A x)) (IS x (All_in A (rec ∘ get_f x))) := 644 | v_trans 645 | (cong (rec_helper _ _ (λ p c ↦ rec _)) 646 | (split_view_isContr A _ _ (split_view_refl A x))) 647 | (cong (IS x) (build_IH_eq A rec x)). 648 | 649 | End induction. 650 | End universe_of_datatypes. -------------------------------------------------------------------------------- /code/General.v: -------------------------------------------------------------------------------- 1 | (* Author: Jasper Hugunin *) 2 | 3 | From WhyNotW Require Import Prelude. 4 | 5 | Section universe_of_datatypes. 6 | Universe i. (* Size of data in our datatypes *) 7 | 8 | (* 9 | We adapt Dybjer-Setzer IR codes to the simple inductive case. 10 | Deviating from the paper, we split infinitary and non-infinitary inductive 11 | arguments to compensate for the lack of strict eta for unit 12 | (without which we don't have (unit → X) equivalent to X). 13 | 14 | We include the choice constructor to avoid getting stuck trying to recurse 15 | under pattern matching on constructor labels. 16 | 17 | Another option would be to take a whole telescope, which would be more uniform 18 | and perhaps more friendly to users. 19 | Since we have strict eta for pairs, that is however overkill, 20 | and for expository purposes the following universe should be just right. 21 | *) 22 | Inductive Code := 23 | | nil 24 | | nonind (A : Type@{i}) (B : A → Code) 25 | | inf_ind (Ix : Type@{i}) (B : Code) 26 | | ind (B : Code) (* morally ind B = inf_ind 1 B *) 27 | | choice (A B : Code) 28 | (* morally choice A B = 29 | nonind 2 (λ b ↦ match b with false => A | true => B end) *) 30 | . 31 | (* Note that Code@{i} has itself a code in Code@{j | i <= j}: 32 | Code@{i} : Code@{j | i <= j} := 33 | nonind ("nil" + "nonind" (A : Type@{i}) + 34 | "inf_ind" (Ix : Type@{i}) + "ind" + "choice") (fun label => 35 | match label with 36 | | "nil" => nil 37 | | "nonind" A => inf_ind A nil 38 | | "inf_ind" Ix => ind nil 39 | | "ind" => ind nil 40 | | "choice" => ind (ind nil) 41 | end) 42 | *) 43 | 44 | (* These codes describe functors: *) 45 | Fixpoint F@{j | i <= j} (A : Code) : Type@{j} → Type@{j} := 46 | match A with 47 | | nil => λ X ↦ 1 48 | | nonind A B => λ X ↦ Σ (a : A), F (B a) X 49 | | inf_ind Ix B => λ X ↦ (Ix → X) × F B X 50 | | ind B => λ X ↦ X × F B X 51 | | choice A B => λ X ↦ F A X ⊔ F B X 52 | end. 53 | 54 | Fixpoint Fmap@{j1 j2} 55 | (A : Code) {X : Type@{j1}} {Y : Type@{j2}} (f : X → Y) 56 | : F@{j1} A X → F@{j2} A Y := 57 | match A with 58 | | nil => λ x ↦ x 59 | | nonind A B => λ x ↦ (x.(fst); Fmap (B x.(fst)) f x.(snd)) 60 | | inf_ind Ix B => λ x ↦ (f ∘ x.(fst); Fmap B f x.(snd)) 61 | | ind B => λ x ↦ (f x.(fst); Fmap B f x.(snd)) 62 | | choice A B => λ x ↦ match x with 63 | | inl a => inl (Fmap A f a) 64 | | inr b => inr (Fmap B f b) 65 | end 66 | end. 67 | 68 | (* Checking lawfulness is left to the reader *) 69 | 70 | (* We next define the standard construction in W types. *) 71 | 72 | Fixpoint shapes (A : Code) : Type@{i} := 73 | match A with 74 | | nil => 1 75 | | nonind A B => Σ (a : A), shapes (B a) 76 | | inf_ind Ix B => shapes B 77 | | ind B => shapes B 78 | | choice A B => shapes A ⊔ shapes B 79 | end. 80 | 81 | Fixpoint positions (A : Code) : shapes A → Type@{i} := 82 | match A with 83 | | nil => λ x ↦ 0 84 | | nonind A B => λ x ↦ positions (B x.(fst)) x.(snd) 85 | | inf_ind Ix B => λ x ↦ Ix ⊔ positions B x 86 | | ind B => λ x ↦ 1 ⊔ positions B x 87 | | choice A B => λ x ↦ match x with 88 | | inl a => positions A a 89 | | inr b => positions B b 90 | end 91 | end. 92 | 93 | Fixpoint get_shape@{j} {A : Code} {X : Type@{j}} : F@{j} A X → shapes A := 94 | match A return F A X → shapes A with 95 | | nil => λ x ↦ x 96 | | nonind A B => λ x ↦ (x.(fst); get_shape x.(snd)) 97 | | inf_ind Ix B => λ x ↦ get_shape x.(snd) 98 | | ind B => λ x ↦ get_shape x.(snd) 99 | | choice A B => λ x ↦ match x with 100 | | inl a => inl (get_shape a) 101 | | inr b => inr (get_shape b) 102 | end 103 | end. 104 | 105 | Fixpoint get_f@{j} {A : Code} {X : Type@{j}} : 106 | ∀ x : F@{j} A X, positions A (get_shape x) → X := 107 | match A with 108 | | nil => λ _ xx ↦ match xx with end 109 | | nonind A B => λ x ↦ get_f x.(snd) 110 | | inf_ind Ix B => λ x p ↦ match p with 111 | | inl i => x.(fst) i 112 | | inr p => get_f x.(snd) p 113 | end 114 | | ind B => λ x p ↦ match p with 115 | | None => x.(fst) 116 | | Some p => get_f x.(snd) p 117 | end 118 | | choice A B => λ x ↦ match x with 119 | | inl a => get_f a 120 | | inr b => get_f b 121 | end 122 | end. 123 | 124 | Definition get_W_arg@{j} {A : Code} {X : Type@{j}} (x : F@{j} A X) : 125 | Σ (s : shapes A), (positions A s → X) := 126 | (get_shape x; get_f x). 127 | 128 | (* These functors induce a refined notion of forall p : positions A s, P p *) 129 | Fixpoint All@{j} (A : Code) : 130 | ∀ {s : shapes A} (P : positions A s → Type@{j}), Type@{j} := 131 | match A with 132 | | nil => λ _ _ ↦ 1 133 | | nonind A B => λ s P ↦ All (B s.(fst)) P 134 | | inf_ind Ix B => λ s P ↦ (∀ i, P (inl i)) × All B (P ∘ inr') 135 | | ind B => λ s P ↦ P None × All B (P ∘ Some') 136 | | choice A B => λ s ↦ match s with 137 | | inl a => All A (s := a) 138 | | inr b => All B (s := b) 139 | end 140 | end. 141 | 142 | Fixpoint All_in@{j} (A : Code) : 143 | ∀ {s : shapes A} {P : positions A s → Type@{j}}, (∀ p, P p) → All@{j} A P := 144 | match A with 145 | | nil => λ s _ _ ↦ s 146 | | nonind A B => λ s P f ↦ All_in (B s.(fst)) f 147 | | inf_ind Ix B => λ s P f ↦ (f ∘ inl'; All_in B (f ∘ inr')) 148 | | ind B => λ s P f ↦ (f None; All_in B (f ∘ Some')) 149 | | choice A B => λ s ↦ match s with 150 | | inl a => λ P f ↦ All_in A (s := a) f 151 | | inr b => λ P f ↦ All_in B (s := b) f 152 | end 153 | end. 154 | 155 | (* We can lift predicates over these functors: *) 156 | Definition liftP@{j k} (A : Code) {X : Type@{j}} (P : X → Type@{k}) 157 | (x : F@{j} A X) : Type@{k} := 158 | All@{k} A (P ∘ get_f x). 159 | 160 | Definition preEl (A : Code) : Type@{i} := W (shapes A) (positions A). 161 | 162 | (* 163 | Then we define the canonical predicate. 164 | A term is canonical if its subterms are canonical 165 | and (s; f) is in the image of get_W_arg. 166 | *) 167 | Definition canonical@{} (A : Code) : preEl A → Type@{i} := 168 | fix canonical x := match x with 169 | | sup s f => 170 | All A (canonical ∘ f) × 171 | Σ (t : F A (preEl A)), Id (get_W_arg t) (s; f) 172 | end. 173 | 174 | Definition El (A : Code) : Type@{i} := Σ (x : preEl A), canonical A x. 175 | 176 | Section split. 177 | (* 178 | We prove an equivalence between F A (Σ (x : X), C x) and 179 | Σ (fx : F A X), All A (C ∘ get_f x) 180 | *) 181 | Context 182 | {X : Type@{i}} 183 | {C : X → Type@{i}} 184 | . 185 | 186 | Fixpoint get_heredity {A : Code} : 187 | ∀ (x : F A (Σ (x : X), C x)), 188 | let x_fst := Fmap A fst x in 189 | All A (s := get_shape x_fst) (C ∘ get_f x_fst) := 190 | match A with 191 | | nil => λ x ↦ x 192 | | nonind A B => λ x ↦ get_heredity x.(snd) 193 | | inf_ind Ix B => λ x ↦ (snd ∘ x.(fst); get_heredity x.(snd)) 194 | | ind B => λ x ↦ (snd x.(fst); get_heredity x.(snd)) 195 | | choice A B => λ x ↦ match x with 196 | | inl a => get_heredity a 197 | | inr b => get_heredity b 198 | end 199 | end. 200 | 201 | (* The pair (Fmap fst, get_heredity) is right-invertible *) 202 | (* 203 | split_view A x hered is equivalent to 204 | Σ (xc : F A (Σ x, C)), Id (Fmap A fst xc; get_hereditary xc) (x; hered) 205 | *) 206 | Inductive split_view (A : Code) : 207 | ∀ (x : F A X), All A (C ∘ get_f x) → Type@{i} := 208 | | split_view_refl (x : F A (Σ (x : X), C x)) : 209 | split_view A (Fmap A fst x) (get_heredity x) 210 | . 211 | Definition split_view_cover_nonind A B x hered : 212 | split_view (B x.(fst)) x.(snd) hered → split_view (nonind A B) x hered := 213 | λ cover ↦ match cover in split_view _ x_snd hered 214 | return split_view (nonind A B) (x.(fst); x_snd) hered 215 | with split_view_refl _ xc => 216 | split_view_refl (nonind A B) (x.(fst); xc) 217 | end. 218 | Definition split_view_cover_inf_ind Ix B x hered : 219 | split_view B x.(snd) hered.(snd) → split_view (inf_ind Ix B) x hered := 220 | λ cover ↦ match cover in split_view _ x_snd hered_snd 221 | return split_view (inf_ind Ix B) (x.(fst); x_snd) (hered.(fst); hered_snd) 222 | with split_view_refl _ xc => 223 | split_view_refl (inf_ind Ix B) (λ i ↦ (x.(fst) i; hered.(fst) i); xc) 224 | end. 225 | Definition split_view_cover_ind B x hered : 226 | split_view B x.(snd) hered.(snd) → split_view (ind B) x hered := 227 | λ cover ↦ match cover in split_view _ x_snd hered_snd 228 | return split_view (ind B) (x.(fst); x_snd) (hered.(fst); hered_snd) 229 | with split_view_refl _ xc => 230 | split_view_refl (ind B) ((x.(fst); hered.(fst)); xc) 231 | end. 232 | Definition split_view_cover_choice_l A B a hered : 233 | split_view A a hered → split_view (choice A B) (inl a) hered := 234 | λ cover ↦ match cover with split_view_refl _ xc => 235 | split_view_refl (choice A B) (inl xc) 236 | end. 237 | Definition split_view_cover_choice_r A B b hered : 238 | split_view B b hered → split_view (choice A B) (inr b) hered := 239 | λ cover ↦ match cover with split_view_refl _ xc => 240 | split_view_refl (choice A B) (inr xc) 241 | end. 242 | Fixpoint split_view_cover (A : Code) : 243 | ∀ x hered, split_view A x hered := 244 | match A with 245 | | nil => λ 'tt 'tt ↦ split_view_refl nil tt 246 | | nonind A B => λ x hered ↦ 247 | split_view_cover_nonind A B _ _ (split_view_cover (B x.(fst)) x.(snd) hered) 248 | | inf_ind Ix B => λ x hered ↦ 249 | split_view_cover_inf_ind Ix B _ _ (split_view_cover B x.(snd) hered.(snd)) 250 | | ind B => λ x hered ↦ 251 | split_view_cover_ind B _ _ (split_view_cover B x.(snd) hered.(snd)) 252 | | choice A B => λ x ↦ match x with 253 | | inl a => λ hered ↦ 254 | split_view_cover_choice_l A B _ _ (split_view_cover A a hered) 255 | | inr b => λ hered ↦ 256 | split_view_cover_choice_r A B _ _ (split_view_cover B b hered) 257 | end 258 | end. 259 | (* And the fibers are in fact contractible *) 260 | Fixpoint split_view_isContr A : 261 | ∀ x hered other, Id other (split_view_cover A x hered) := 262 | match A with 263 | | nil => λ x hered '(split_view_refl _ tt) ↦ refl 264 | | nonind A B => λ x hered '(split_view_refl _ xc) ↦ 265 | cong (split_view_cover_nonind A B (Fmap _ fst xc) (get_heredity xc)) 266 | (split_view_isContr _ _ _ (split_view_refl _ xc.(snd))) 267 | | inf_ind Ix B => λ x hered '(split_view_refl _ xc) ↦ 268 | cong (split_view_cover_inf_ind Ix B (Fmap _ fst xc) (get_heredity xc)) 269 | (split_view_isContr _ _ _ (split_view_refl _ xc.(snd))) 270 | | ind B => λ x hered '(split_view_refl _ xc) ↦ 271 | cong (split_view_cover_ind B (Fmap _ fst xc) (get_heredity xc)) 272 | (split_view_isContr _ _ _ (split_view_refl _ xc.(snd))) 273 | | choice A B => λ x hered '(split_view_refl _ xc) ↦ match xc with 274 | | inl a => 275 | cong (split_view_cover_choice_l A B (Fmap _ fst a) (get_heredity a)) 276 | (split_view_isContr _ _ _ (split_view_refl _ a)) 277 | | inr b => 278 | cong (split_view_cover_choice_r A B (Fmap _ fst b) (get_heredity b)) 279 | (split_view_isContr _ _ _ (split_view_refl _ b)) 280 | end 281 | end. 282 | End split. 283 | 284 | Definition intro (A : Code) : F@{i} A (El A) → El A := 285 | λ x ↦ let pre_x := Fmap A fst x in 286 | (sup (get_shape pre_x) (get_f pre_x); get_heredity x; pre_x; refl). 287 | 288 | (* Now for induction *) 289 | 290 | Section induction. 291 | Universe j. 292 | Context 293 | (A : Code) 294 | (P : El A → Type@{j}) 295 | (IS : ∀ (x : F A (El A)), liftP@{i j} A P x → P (intro A x)) 296 | . 297 | 298 | Fixpoint build_IH@{} (A' : Code) : 299 | ∀ x (IH : ∀ p c, P (get_f (Fmap@{i i} A' fst x) p; c)), 300 | All@{j} A' (P ∘ get_f x) := 301 | match A' with 302 | | nil => λ x IH ↦ x 303 | | nonind A B => λ x IH ↦ build_IH (B x.(fst)) x.(snd) IH 304 | | inf_ind Ix B => λ x IH ↦ 305 | (λ i ↦ IH (inl i) (x.(fst) i).(snd); build_IH B x.(snd) (IH ∘ inr')) 306 | | ind B => λ x IH ↦ 307 | (IH None x.(fst).(snd); build_IH B x.(snd) (IH ∘ Some')) 308 | | choice A B => λ x ↦ match x with 309 | | inl a => build_IH A a 310 | | inr b => build_IH B b 311 | end 312 | end. 313 | 314 | Fixpoint build_IH_eq@{} (A' : Code) (el : ∀ x, P x) : 315 | ∀ x, 316 | Id (build_IH A' x (λ p c ↦ el _)) (All_in A' (λ p ↦ el _)) := 317 | match A' with 318 | | nil => λ x ↦ refl 319 | | nonind A B => λ x ↦ build_IH_eq (B x.(fst)) el x.(snd) 320 | | inf_ind Ix B => λ x ↦ 321 | cong (pair (el ∘ x.(fst))) (build_IH_eq B el x.(snd)) 322 | | ind B => λ x ↦ 323 | cong (pair (el x.(fst))) (build_IH_eq B el x.(snd)) 324 | | choice A B => λ x ↦ match x with 325 | | inl a => build_IH_eq A el a 326 | | inr b => build_IH_eq B el b 327 | end 328 | end. 329 | 330 | Definition rec_helper@{} t (hered : All A (canonical A ∘ _)) 331 | (IH : ∀ p c, P (get_f t p; c)) : 332 | split_view A t hered → P (sup _ _; hered; t; refl) := 333 | fun cover => match cover with split_view_refl _ x => λ IH ↦ 334 | IS x (build_IH A x IH) 335 | end IH. 336 | Definition rec@{} : ∀ x, P x := 337 | let fix rec (x : preEl A) : ∀ c, P (x; c) := match x with 338 | | sup s f => λ '((hered; t; eq) : canonical A (sup s f)) ↦ 339 | match eq in Id _ (s; f) 340 | return 341 | ∀ hered : All A (canonical A ∘ f), 342 | (∀ p c, P (f p; c)) → 343 | P (sup s f; hered; t; eq) 344 | with refl => λ hered IH ↦ 345 | rec_helper t hered IH (split_view_cover A t hered) 346 | end hered (λ p ↦ rec (f p)) 347 | end in 348 | λ x ↦ rec x.(fst) x.(snd). 349 | 350 | Definition rec_eq@{} x 351 | : Id (rec (intro A x)) (IS x (All_in A (rec ∘ get_f x))) := 352 | v_trans 353 | (cong (rec_helper _ _ (λ p c ↦ rec _)) 354 | (split_view_isContr A _ _ (split_view_refl A x))) 355 | (cong (IS x) (build_IH_eq A rec x)). 356 | 357 | End induction. 358 | End universe_of_datatypes. -------------------------------------------------------------------------------- /code/Makefile: -------------------------------------------------------------------------------- 1 | all: Makefile.coq 2 | @+$(MAKE) -f Makefile.coq all 3 | 4 | clean: Makefile.coq 5 | @+$(MAKE) -f Makefile.coq cleanall 6 | @rm -f Makefile.coq Makefile.coq.conf 7 | 8 | Makefile.coq: _CoqProject 9 | $(COQBIN)coq_makefile -f _CoqProject -o Makefile.coq 10 | 11 | force _CoqProject Makefile: ; 12 | 13 | %: Makefile.coq force 14 | @+$(MAKE) -f Makefile.coq $@ 15 | 16 | .PHONY: all clean force 17 | -------------------------------------------------------------------------------- /code/Nat.v: -------------------------------------------------------------------------------- 1 | (* Author: Jasper Hugunin 2 | 3 | Here we prove that W types plus Id types with strict J and Sg types with eta 4 | are enough to define N without using function extensionality. 5 | 6 | Note that function extensionality breaks canonicity for this type 7 | (by breaking canonicity for Id). However, as long as Id satisfies J strictly, 8 | this still computes, which is better than the standard encoding. 9 | 10 | Working in Coq we lack strict eta for unit, thus have to work a little harder 11 | in the successor case than in the paper, but it proves no obstacle. 12 | (* *) 13 | *) 14 | From Coq Require Import Init.Ltac. 15 | 16 | From WhyNotW Require Import Prelude. 17 | 18 | Notation OO := false. 19 | Notation SS := true. 20 | 21 | (* Here is the definition of N: *) 22 | Definition nat_pre : Set := 23 | W@{Set} 2 (λ b ↦ match b with OO => 0 | SS => 1 end). 24 | Fixpoint canonical (n : nat_pre) : Set 25 | := match n with 26 | | sup OO f => 27 | Id (λ XX : 0 ↦ match XX return nat_pre with end) f 28 | | sup SS f => 29 | Σ (prev : nat_pre), 30 | Id (λ _ ↦ prev) f × canonical (f ★) 31 | end. 32 | 33 | Definition nat : Set := Σ (n : nat_pre), canonical n. 34 | Definition O : nat := (sup OO _; refl). 35 | Definition S (n : nat) : nat := (sup SS _; fst n; refl; snd n). 36 | 37 | (* 38 | And here we define induction, 39 | first with an explicit proof term and then again by tactics. 40 | Read whichever you are most comfortable with. 41 | *) 42 | Section induction. 43 | Universe i. 44 | Context 45 | (P : nat → Type@{i}) 46 | (ISO : P O) 47 | (ISS : ∀ n, P n → P (S n)) 48 | . 49 | Definition nat_rect@{} : ∀ n, P n 50 | := let fix nat_rect (n_pre : nat_pre) 51 | : ∀ n_canon : canonical n_pre, P (pair n_pre n_canon) 52 | := match n_pre with 53 | | sup OO f => λ 'refl ↦ ISO 54 | | sup SS f => 55 | λ '(pair pred (pair eq pred_canon) : canonical (sup true f)) ↦ 56 | match eq in Id _ f 57 | return 58 | ∀ pred_canon : canonical (f ★), 59 | P (pair (f ★) pred_canon) → 60 | P (pair (sup true f) (pair pred (pair eq pred_canon))) 61 | with refl => λ pred_canon IH ↦ ISS (pair pred pred_canon) IH end 62 | pred_canon (nat_rect (f ★) pred_canon) 63 | end in 64 | λ '(pair pre canon) ↦ nat_rect pre canon. 65 | 66 | Definition nat_rect_ltac@{} : ∀ n, P n. 67 | intros [n_pre n_canon]. 68 | induction n_pre as [a f IH] using W_rect@{Set i}; destruct a. 69 | - (* case S *) 70 | destruct n_canon as [pred [eq pred_canon]]. 71 | destruct eq. 72 | apply (ISS (pair pred pred_canon)). 73 | apply (IH ★). 74 | - (* case O *) 75 | simpl in n_canon; rename n_canon into eq. 76 | destruct eq. 77 | apply ISO. 78 | Defined. 79 | 80 | (* And then we check that the expected equations hold strictly. *) 81 | Definition test_eq_O := convertible (nat_rect O) ISO. 82 | Definition test_eq_S n := convertible (nat_rect (S n)) (ISS n (nat_rect n)). 83 | 84 | Definition test_eq_O' := convertible (nat_rect_ltac O) ISO. 85 | Definition test_eq_S' n := 86 | convertible (nat_rect_ltac (S n)) (ISS n (nat_rect_ltac n)). 87 | 88 | End induction. -------------------------------------------------------------------------------- /code/Prelude.v: -------------------------------------------------------------------------------- 1 | (* Author: Jasper Hugunin *) 2 | 3 | (** 4 | We use the -noinit option in _CoqProject to start from a blank slate. 5 | We also use -indices-matter, to be better compatible with HoTT. 6 | *) 7 | 8 | #[export] Set Universe Polymorphism. 9 | (* #[export] Set Polymorphic Inductive Cumulativity. (* not needed, just cool *) *) 10 | #[export] Set Primitive Projections. (* strict eta for pairs *) 11 | #[export] Unset Elimination Schemes. (* we'll declare only the ones we need *) 12 | 13 | (* 14 | We'll avoid conflicting with Coq's notation levels. 15 | Consists of just Reserved Notations and Scopes. 16 | *) 17 | From Coq Require Export Init.Notations. 18 | 19 | (** We use Unicode characters freely in our notations. *) 20 | 21 | (** * Functions *) 22 | 23 | Notation "'Π' x .. y , P" := (forall x, .. (forall y, P) ..) 24 | (at level 200, x binder, y binder, right associativity, 25 | format "'[ ' '[ ' Π x .. y ']' , '/' P ']'") : type_scope. 26 | Notation "∀ x .. y , P" := (forall x, .. (forall y, P) ..) 27 | (at level 200, x binder, y binder, right associativity, 28 | format "'[ ' '[ ' ∀ x .. y ']' , '/' P ']'") : type_scope. 29 | 30 | (** Can't go far without implication. *) 31 | Notation "A -> B" := (forall (_ : A), B) : type_scope. 32 | Notation "A → B" := (A -> B) 33 | (at level 99, right associativity, B at level 200) : type_scope. 34 | 35 | Notation "'λ' x .. y ↦ z" := (fun x => .. (fun y => z) ..) 36 | (at level 8, x binder, y binder, z at level 200, right associativity) 37 | : core_scope. 38 | Notation "'λ' x ↦ y" := (fun x => y) 39 | (* Hack to mostly hide types of function binders in printing. *) 40 | (at level 0, x ident, y at level 200, right associativity, only printing) 41 | : core_scope. 42 | 43 | (** composition of functions *) 44 | Notation "f ∘ g" := (λ x ↦ f (g x)) 45 | (at level 40, left associativity) : core_scope. 46 | 47 | (** * Finite types *) 48 | Inductive void : Set := . 49 | Notation "⊥" := void : type_scope. 50 | Notation "0" := ⊥ : type_scope. 51 | 52 | Inductive unit : Set := tt. 53 | Notation "⊤" := unit : type_scope. 54 | Notation "1" := ⊤ : type_scope. 55 | Notation "★" := tt : core_scope. 56 | 57 | Inductive bool : Set := true | false. 58 | Notation "2" := bool : type_scope. 59 | 60 | Notation "( b ? x || y )" := 61 | match b as b' return _ with true => x | false => y end 62 | (only printing) : core_scope. 63 | Notation "( b ? x || y )" := 64 | match b with true => x | false => y end 65 | : core_scope. 66 | 67 | Notation "! x" := match x with end 68 | (at level 35, x constr, right associativity, only printing) : core_scope. 69 | Notation "! x" := match x as x' return _ with end 70 | (at level 35, right associativity) : core_scope. 71 | 72 | (** * Dependent pairs *) 73 | Record prod@{i} (A : Type@{i}) (B : A → Type@{i}) : Type@{i} := 74 | pair { fst : A ; snd : B fst }. 75 | Arguments pair {A B} & fst snd. 76 | Arguments fst {A B} _. 77 | Arguments snd {A B} _. 78 | 79 | Notation "x .1" := (x.(fst)) 80 | (at level 1, left associativity, format "x .1") : core_scope. 81 | Notation "x .2" := (x.(snd)) 82 | (at level 1, left associativity, format "x .2") : core_scope. 83 | 84 | Notation "∃ x .. y , P" := (prod _ (λ x ↦ .. (prod _ (λ y ↦ P)) ..)) 85 | (at level 200, x binder, y binder, right associativity, 86 | format "'[ ' '[ ' ∃ x .. y ']' , '/' P ']'") 87 | : type_scope. 88 | Notation "'Σ' x .. y , P" := (prod _ (λ x ↦ .. (prod _ (λ y ↦ P)) ..)) 89 | (at level 200, x binder, y binder, right associativity, 90 | format "'[ ' '[ ' Σ x .. y ']' , '/' P ']'") 91 | : type_scope. 92 | Notation "( x , .. , y , z )" := (pair x .. (pair y z) ..) : core_scope. 93 | Notation "( x ; .. ; y ; z )" := (pair x .. (pair y z) ..) : core_scope. 94 | 95 | Notation "A * B" := (prod A (λ _ ↦ B)) : type_scope. 96 | Notation "A ∧ B" := (prod A (λ _ ↦ B)) 97 | (at level 80, right associativity) : type_scope. 98 | Notation "A × B" := (prod A (λ _ ↦ B)) 99 | (at level 80, right associativity) : type_scope. 100 | 101 | (* tuple notation *) 102 | Notation "[ x ; .. ; y ]" := (pair x .. (pair y ★) ..) : core_scope. 103 | Notation "[ ]" := ★ : core_scope. 104 | 105 | (** * W types *) 106 | Inductive W@{i} (A : Type@{i}) (B : A → Type@{i}) : Type@{i} := 107 | | sup (a : A) (f : B a → W A B) : W A B. 108 | Arguments sup {A B} & a f. 109 | 110 | Definition W_rect@{i j} (A : Type@{i}) (B : A → Type@{i}) (P : W A B → Type@{j}) 111 | (IS : ∀ a f, (∀ b, P (f b)) → P (sup a f)) : ∀ w, P w := 112 | fix rect w := match w with sup a f => IS a f (rect ∘ f) end. 113 | 114 | (** * Identity/Path types *) 115 | Inductive Id@{i} (A : Type@{i}) (x : A) : A → Type@{i} := refl : Id A x x. 116 | Arguments Id {A} x y. 117 | Arguments refl {A} x, {A x}. 118 | 119 | (** A few lemmas about Id *) 120 | Definition cong@{i j} {A : Type@{i}} {B : Type@{j}} (f : A → B) {x y : A} 121 | (e : Id x y) : Id (f x) (f y) := 122 | match e with refl => refl end. 123 | (* Transitivity from the middle. *) 124 | Definition v_trans@{i} {A : Type@{i}} {x y z : A} 125 | (e1 : Id y x) (e2 : Id y z) : Id x z := 126 | match e1 with refl => e2 end. 127 | 128 | (** * Sum types *) 129 | (** We define sum types from bool and Σ *) 130 | 131 | Notation sum A B := 132 | (Σ (b : 2), match b with false => A | true => B end). 133 | Definition sum'@{i} (A B : Type@{i}) : Type@{i} := sum A B. 134 | Notation inl a := (false; a). 135 | Notation inr b := (true; b). 136 | Definition inl'@{i} {A B : Type@{i}} (a : A) : sum'@{i} A B := inl a. 137 | Definition inr'@{i} {A B : Type@{i}} (b : B) : sum'@{i} A B := inr b. 138 | 139 | Notation "A + B" := (sum A B) : type_scope. 140 | Notation "A ∨ B" := (sum A B) (at level 85, right associativity) : type_scope. 141 | Notation "A ⊎ B" := (sum A B) (at level 85, right associativity) : type_scope. 142 | Notation "A ⊔ B" := (sum A B) (at level 85, right associativity) : type_scope. 143 | 144 | Definition option@{i} (A : Type@{i}) : Type@{i} := sum 1 A. 145 | Notation None := (inl ★). 146 | Notation Some a := (inr a). 147 | Definition Some'@{i} {A : Type@{i}} (a : A) : option@{i} A := Some a. 148 | 149 | (** * Universe *) 150 | Definition U@{i} := Type@{i}. 151 | 152 | (** * Notation for tests *) 153 | Notation convertible x y := (refl : Id x y). 154 | -------------------------------------------------------------------------------- /code/README.md: -------------------------------------------------------------------------------- 1 | To compile, simply run `make` in this the `code` subdirectory with Coq (8.12) on your path or with environment variable `COQBIN` set. This may take an hour or two to compile, due to some slow files dealing with bootstrapping. 2 | 3 | * `Nat.v` contains the special case of the construction of the natural numbers 4 | * `General.v` works out the machinery for the general case 5 | * `TestGeneral.v` applies the general construction to several specific inductives 6 | and checks that we get the expected computational behavior. 7 | This includes the type `Code` itself, demonstrating that bootstrapping should be possible. 8 | * `Bootstrap.v` defines `Code` explicitly in terms of `W` types, and then defines the rest of the general 9 | construction on top. This is much slower than `General.v`, taking about half an hour to compile on my machine. 10 | `TestBootstrap/Nat.v` applies the general construction in `Bootstrap.v` to the natural numbers 11 | (about 20 minutes to compile). `TestBootstrap/TiedKnot.v` checks that `Code = El "Code"`. 12 | -------------------------------------------------------------------------------- /code/TestBootstrap/Nat.v: -------------------------------------------------------------------------------- 1 | (* Author: Jasper Hugunin *) 2 | 3 | From WhyNotW Require Import Prelude. 4 | From WhyNotW Require Import Bootstrap. 5 | 6 | (** * Natural numbers: Inductive nat := O | S (n : nat). *) 7 | 8 | (* Constructor tags *) 9 | Notation O' x := (false; x). 10 | Notation S' x := (true; x). 11 | 12 | Definition nat_code : Code@{Set _} := choice nil (ind nil). 13 | 14 | Definition nat := El nat_code. 15 | Definition O := intro nat_code (O' []). 16 | Definition S n := intro nat_code (S' [n]). 17 | 18 | Section nat_rect. 19 | Universe j. 20 | Context 21 | (P : nat → Type@{j}) 22 | (ISO : P O) 23 | (ISS : ∀ n, P n → P (S n)) 24 | . 25 | Definition IS arg : liftP@{Set _ Set j} nat_code P arg → P (intro nat_code arg) := 26 | match arg with 27 | | O' [] => λ '[] ↦ ISO 28 | | S' [n] => λ '[IH] ↦ ISS n IH 29 | end. 30 | Definition nat_rect : ∀ n, P n := rec nat_code P IS. 31 | 32 | (* Check that the expected equations hold by rec_eq = refl *) 33 | Definition test_eq_O := convertible 34 | (rec_eq nat_code P IS (O' [])) 35 | (convertible (nat_rect O) ISO). 36 | Definition test_eq_S n := convertible 37 | (rec_eq nat_code P IS (S' [n])) 38 | (convertible (nat_rect (S n)) (ISS n (nat_rect n))). 39 | 40 | End nat_rect. -------------------------------------------------------------------------------- /code/TestBootstrap/TiedKnot.v: -------------------------------------------------------------------------------- 1 | (* Author: Jasper Hugunin *) 2 | 3 | From WhyNotW Require Import Prelude. 4 | From WhyNotW Require Import Bootstrap. 5 | From Coq Require Import Ltac. 6 | 7 | (* We see that Code = El "Code". *) 8 | 9 | Section test_code. 10 | Universes i i' i''. 11 | Constraint i < i', i' < i''. 12 | 13 | Definition Code_code : Code@{i' i''} := 14 | choice 15 | (choice 16 | nil (* nil' *) 17 | (choice 18 | (nonind Type@{i} λ A ↦ inf_ind A nil) (* nonind' *) 19 | (ind (ind nil)))) (* choice' *) 20 | (choice 21 | (nonind Type@{i} λ Ix ↦ ind nil) (* inf_ind' *) 22 | (ind nil)). (* ind' *) 23 | 24 | (* Code codes for itself *) 25 | Definition test_bootstrap := convertible 26 | Code@{i i'} 27 | (El@{i' i''} Code_code). 28 | 29 | End test_code. -------------------------------------------------------------------------------- /code/TestGeneral.v: -------------------------------------------------------------------------------- 1 | (* Author: Jasper Hugunin *) 2 | 3 | From WhyNotW Require Import Prelude. 4 | From WhyNotW Require Import General. 5 | 6 | (** 7 | We specialize the general case to several specific inductive types, 8 | and check that we get the expected behavior. 9 | *) 10 | 11 | Module test_nat. 12 | (** * Natural numbers: Inductive nat := O | S (n : nat). *) 13 | 14 | (* Constructor tags *) 15 | Notation O' x := (false; x). 16 | Notation S' x := (true; x). 17 | 18 | Definition nat_code : Code@{Set} := choice nil (ind nil). 19 | 20 | Definition nat := El nat_code. 21 | Definition O := intro nat_code (O' []). 22 | Definition S n := intro nat_code (S' [n]). 23 | 24 | Section nat_rect. 25 | Universe j. 26 | Context 27 | (P : nat → Type@{j}) 28 | (ISO : P O) 29 | (ISS : ∀ n, P n → P (S n)) 30 | . 31 | Definition IS arg : liftP@{Set Set j} nat_code P arg → P (intro nat_code arg) := 32 | match arg with 33 | | O' [] => λ '[] ↦ ISO 34 | | S' [n] => λ '[IH] ↦ ISS n IH 35 | end. 36 | Definition nat_rect : ∀ n, P n := rec nat_code P IS. 37 | 38 | (* Check that the expected equations hold by rec_eq = refl *) 39 | Definition test_eq_O := convertible 40 | (rec_eq nat_code P IS (O' [])) 41 | (convertible (nat_rect O) ISO). 42 | Definition test_eq_S n := convertible 43 | (rec_eq nat_code P IS (S' [n])) 44 | (convertible (nat_rect (S n)) (ISS n (nat_rect n))). 45 | 46 | End nat_rect. 47 | End test_nat. 48 | 49 | Module test_list. 50 | (** * Lists: Inductive list A := nil | cons (a : A) (l : list A). *) 51 | (* Because we are working internally, uniform parameters are free. *) 52 | Section test_list. 53 | Universe i. 54 | Context (A : Type@{i}). 55 | 56 | (* Constructor tags *) 57 | Notation nil' x := (false; x). 58 | Notation cons' x := (true; x). 59 | 60 | Definition list_code : Code@{i} := 61 | choice nil (nonind A λ a ↦ ind nil). 62 | 63 | Definition list : Type@{i} := El list_code. 64 | Definition nil : list := intro list_code (nil' []). 65 | Definition cons (a : A) (l : list) : list := intro list_code (cons' [a; l]). 66 | 67 | Section list_rect. 68 | Universe j. 69 | Context 70 | (P : list -> Type@{j}) 71 | (IS_nil : P nil) 72 | (IS_cons : ∀ a l, P l → P (cons a l)) 73 | . 74 | 75 | Definition IS arg : liftP@{i i j} list_code P arg → P (intro list_code arg) := 76 | match arg with 77 | | nil' [] => λ '[] ↦ IS_nil 78 | | cons' [a; l] => λ '[IH] ↦ IS_cons a l IH 79 | end. 80 | Definition list_rect : ∀ l, P l := rec list_code P IS. 81 | 82 | (* Check that the expected equations hold by rec_eq = refl *) 83 | Definition test_eq_nil := convertible 84 | (rec_eq list_code P IS (nil' [])) 85 | (convertible (list_rect nil) IS_nil). 86 | Definition test_eq_cons a l := convertible 87 | (rec_eq list_code P IS (cons' [a; l])) 88 | (convertible (list_rect (cons a l)) (IS_cons a l (list_rect l))). 89 | 90 | End list_rect. 91 | End test_list. 92 | End test_list. 93 | 94 | Module test_code. 95 | (** * The type of codes themselves: this can be used to bootstrap. *) 96 | (** This tests infinitary arguments *) 97 | (* 98 | Inductive Code := 99 | | nil 100 | | nonind (A : Type@{i}) (B : A → Code) 101 | | choice (A B : Code) 102 | | inf_ind (Ix : Type@{i}) (B : Code) 103 | | ind (B : Code) (* morally ind B = inf_ind 1 B *) 104 | . 105 | *) 106 | (* Constructor tags *) 107 | Notation nil' x := (false; false; x). 108 | Notation nonind' x := (false; true; false; x). 109 | Notation choice' x := (false; true; true; x). 110 | Notation inf_ind' x := (true; false; x). 111 | Notation ind' x := (true; true; x). 112 | 113 | Section test_code. 114 | Universe i. 115 | Universe i'. 116 | Constraint i < i'. 117 | 118 | Definition Code_code : Code@{i'} := 119 | choice 120 | (choice 121 | nil (* nil' *) 122 | (choice 123 | (nonind Type@{i} λ A ↦ inf_ind A nil) (* nonind' *) 124 | (ind (ind nil)))) (* choice' *) 125 | (choice 126 | (nonind Type@{i} λ Ix ↦ ind nil) (* inf_ind' *) 127 | (ind nil)). (* ind' *) 128 | 129 | (* Note that Code@{i} is in Type@{i+1}. *) 130 | Definition Code_ : Type@{i'} := El Code_code. 131 | Definition nil_ : Code_ := intro Code_code (nil' []). 132 | Definition nonind_ (A : Type@{i}) (B : A → Code_) : Code_ := 133 | intro Code_code (nonind' [A; B]). 134 | Definition inf_ind_ (Ix : Type@{i}) (B : Code_) : Code_ := 135 | intro Code_code (inf_ind' [Ix; B]). 136 | Definition ind_ (B : Code_) : Code_ := 137 | intro Code_code (ind' [B]). 138 | Definition choice_ (A B : Code_) : Code_ := 139 | intro Code_code (choice' [A; B]). 140 | 141 | Section code_rect. 142 | Universe j. 143 | Constraint i' <= j. 144 | Context 145 | (P : Code_ → Type@{j}) 146 | (IS_nil : P nil_) 147 | (IS_nonind : ∀ (A : Type@{i}) B, (∀ a, P (B a)) → P (nonind_ A B)) 148 | (IS_inf_ind : ∀ (Ix : Type@{i}) B, P B → P (inf_ind_ Ix B)) 149 | (IS_ind : ∀ B, P B → P (ind_ B)) 150 | (IS_choice : ∀ A B, P A → P B → P (choice_ A B)) 151 | . 152 | 153 | Definition IS arg : liftP@{i' i' j} Code_code P arg → P (intro Code_code arg) := 154 | match arg with 155 | | nil' [] => λ '[] ↦ IS_nil 156 | | nonind' [A; B] => λ '[IH] ↦ IS_nonind A B IH 157 | | inf_ind' [Ix; B] => λ '[IH] ↦ IS_inf_ind Ix B IH 158 | | ind' [B] => λ '[IH] ↦ IS_ind B IH 159 | | choice' [A; B] => λ '[IHA; IHB] ↦ IS_choice A B IHA IHB 160 | end. 161 | Definition Code_rect@{|} : ∀ A, P A := rec Code_code P IS. 162 | 163 | (* Check that the expected equations hold by rec_eq = refl *) 164 | Definition test_eq_nil := convertible 165 | (rec_eq Code_code P IS (nil' [])) 166 | (convertible (Code_rect nil_) IS_nil). 167 | Definition test_eq_nonind A B := convertible 168 | (rec_eq Code_code P IS (nonind' [A; B])) 169 | (convertible 170 | (Code_rect (nonind_ A B)) 171 | (IS_nonind A B (Code_rect ∘ B))). 172 | Definition test_eq_inf_ind Ix B := convertible 173 | (rec_eq Code_code P IS (inf_ind' [Ix; B])) 174 | (convertible (Code_rect (inf_ind_ Ix B)) (IS_inf_ind Ix B (Code_rect B))). 175 | Definition test_eq_ind B := convertible 176 | (rec_eq Code_code P IS (ind' [B])) 177 | (convertible (Code_rect (ind_ B)) (IS_ind B (Code_rect B))). 178 | Definition test_eq_choice A B := convertible 179 | (rec_eq Code_code P IS (choice' [A; B])) 180 | (convertible 181 | (Code_rect (choice_ A B)) 182 | (IS_choice A B (Code_rect A) (Code_rect B))). 183 | End code_rect. 184 | End test_code. 185 | 186 | Module bootstrap. 187 | (* Code_ doesn't depend on any inductive except W *) 188 | Definition sum_rect@{i j} {A B : Type@{i}} {P : sum@{i} A B → Type@{j}} 189 | (f : ∀ a, P (inl a)) (g : ∀ b, P (inr b)) s : P s := 190 | match s with 191 | | inl a => f a 192 | | inr b => g b 193 | end. 194 | Definition test_Code_eq@{i i' +} := convertible Code_@{i i'} 195 | (let shapes := (1 ⊔ Type@{i} × 1 ⊔ 1) ⊔ Type@{i} × 1 ⊔ 1 in 196 | let positions : shapes → Type@{i'} := 197 | sum_rect 198 | (sum_rect 199 | (λ _ ↦ 0) (* nil *) 200 | (sum_rect 201 | (λ (args : Type@{i} × 1) ↦ args.(fst) ⊔ 0) (* nonind *) 202 | (λ _ ↦ 1 ⊔ (1 ⊔ 0)))) (* choice *) 203 | (sum_rect 204 | (λ _ ↦ 1 ⊔ 0) (* inf_ind *) 205 | (λ _ ↦ 1 ⊔ 0)) in 206 | let preCode := W shapes positions in 207 | let _ := convertible preCode (preEl Code_code) in 208 | Σ (pre_code : (W@{i'} shapes positions)), (* ind *) 209 | let fix canonical x : Type@{i'} := match x with 210 | | sup a f => 211 | (sum_rect (P := λ a ↦ (positions a → Type@{i'}) → Type@{i'}) 212 | (sum_rect (P := λ a ↦ (positions (inl a) → Type@{i'}) → Type@{i'}) 213 | (λ _ _ ↦ 1) 214 | (sum_rect (P := λ a ↦ (positions (inl (inr a)) → Type@{i'}) → Type@{i'}) 215 | (λ args canon ↦ (∀ i, canon (inl i)) × 1) 216 | (λ args canon ↦ canon None × canon (Some None) × 1))) 217 | (sum_rect (P := λ a ↦ (positions (inr a) → Type@{i'}) → Type@{i'}) 218 | (λ args canon ↦ canon None × 1) 219 | (λ args canon ↦ canon None × 1)) 220 | a (canonical ∘ f)) × 221 | Σ (t : 222 | (1 ⊔ (* nil *) 223 | (Σ (A : Type@{i}), (A → preCode) × 1) ⊔ (* nonind *) 224 | (preCode × preCode × 1) (* choice *)) ⊔ 225 | (Type@{i} × preCode × 1) (* inf_ind *) ⊔ 226 | (preCode × 1) (* ind *)), 227 | let get_shape : _ → shapes := 228 | sum_rect (P := λ _ ↦ shapes) 229 | (λ t ↦ inl (sum_rect (P := λ _ ↦ 1 ⊔ Type@{i} × 1 ⊔ 1) 230 | (λ args ↦ inl args) (* nil *) 231 | (λ t ↦ inr (sum_rect (P := λ _ ↦ Type@{i} × 1 ⊔ 1) 232 | (λ (args : _ × _ × 1) ↦ inl (args.(fst); args.(snd).(snd))) (* nonind *) 233 | (λ (args : _ × _ × 1) ↦ inr args.(snd).(snd)) (* choice *) 234 | t)) 235 | t)) 236 | (λ t ↦ inr (sum_rect (P := λ _ ↦ Type@{i} × 1 ⊔ 1) 237 | (λ args : _ × _ × 1 ↦ inl (args.(fst); args.(snd).(snd))) (* inf_ind *) 238 | (λ args ↦ inr args.(snd)) (* ind *) 239 | t)) in 240 | let a' : shapes := get_shape t in 241 | let f' : positions a' → preCode := 242 | sum_rect (P := λ t ↦ positions (get_shape t) → preCode) 243 | (sum_rect (P := λ t ↦ positions (get_shape (inl t)) → preCode) 244 | (λ _ (XX : 0) ↦ match XX with end) (* nil *) 245 | (sum_rect (P := λ t ↦ positions (get_shape (inl (inr t))) → preCode) 246 | (λ args ↦ sum_rect (* nonind *) 247 | (args.(snd).(fst)) 248 | (λ XX : 0 ↦ match XX with end)) 249 | (λ args ↦ sum_rect (* choice *) 250 | (λ '★ ↦ args.(fst)) 251 | (sum_rect 252 | (λ '★ ↦ args.(snd).(fst)) 253 | (λ XX : 0 ↦ match XX with end))))) 254 | (sum_rect (P := λ t ↦ positions (get_shape (inr t)) → preCode) 255 | (λ args ↦ sum_rect (* inf_ind *) 256 | (λ '★ ↦ args.(snd).(fst)) 257 | (λ XX : 0 ↦ match XX with end)) 258 | (λ args ↦ sum_rect (* ind *) 259 | (λ '★ ↦ args.(fst)) 260 | (λ XX : 0 ↦ match XX with end))) 261 | t in 262 | Id (a'; f') (a; f) 263 | end in 264 | canonical pre_code). 265 | (* 266 | Code_rect also doesn't depend on Code when normalized, thus we can 267 | go back to the top of General.v and bootstrap using the normalized definitions. 268 | 269 | After doing so, we have Code@{i} = El Code_code, with Code_code : Code@{i+1}. 270 | *) 271 | End bootstrap. 272 | End test_code. 273 | -------------------------------------------------------------------------------- /code/_CoqProject: -------------------------------------------------------------------------------- 1 | -Q . WhyNotW 2 | -arg -w -arg +all 3 | -arg -noinit 4 | -arg -indices-matter 5 | Prelude.v 6 | Nat.v 7 | General.v 8 | TestGeneral.v 9 | Bootstrap.v 10 | TestBootstrap/Nat.v 11 | TestBootstrap/TiedKnot.v 12 | -------------------------------------------------------------------------------- /paper/cc-by.pdf: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/jashug/WhyNotW/c1f26bde6f45519253ed766b99136649eaf9b5bb/paper/cc-by.pdf -------------------------------------------------------------------------------- /paper/lipics-logo-bw.pdf: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/jashug/WhyNotW/c1f26bde6f45519253ed766b99136649eaf9b5bb/paper/lipics-logo-bw.pdf -------------------------------------------------------------------------------- /paper/lipics-v2019.cls: -------------------------------------------------------------------------------- 1 | %% 2 | %% This is file `lipics-v2019.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) 2018 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-v2019} 27 | [2020/04/29 v2.2.1 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\copyrightline{% 134 | \ifx\@hideLIPIcs\@undefined 135 | \ifx\@EventLogo\@empty 136 | \else 137 | \setbox\@tempboxa\hbox{\includegraphics[height=42\p@]{\@EventLogo}}% 138 | \rlap{\hspace\textwidth\hspace{-\wd\@tempboxa}\hspace{\z@}% 139 | \vtop to\z@{\vskip-0mm\unhbox\@tempboxa\vss}}% 140 | \fi 141 | \scriptsize 142 | \vtop{\hsize\textwidth 143 | \nobreakspace\par 144 | \@Copyright 145 | \ifx\@EventLongTitle\@empty\else\@EventLongTitle.\\\fi 146 | \ifx\@EventEditors\@empty\else 147 | \@Eds: \@EventEditors 148 | ; Article~No.\,\@ArticleNo; pp.\,\@ArticleNo:\thepage--\@ArticleNo:\pageref{LastPage}% 149 | \\ 150 | \fi 151 | \setbox\@tempboxa\hbox{\includegraphics[height=14\p@,trim=0 15 0 0]{lipics-logo-bw}}% 152 | \hspace*{\wd\@tempboxa}\enskip 153 | \href{https://www.dagstuhl.de/lipics/}% 154 | {Leibniz International Proceedings in Informatics}\\ 155 | \smash{\unhbox\@tempboxa}\enskip 156 | \href{https://www.dagstuhl.de}% 157 | {Schloss Dagstuhl -- Leibniz-Zentrum f{\"u}r Informatik, Dagstuhl Publishing, Germany}}% 158 | \fi} 159 | \def\ps@plain{\let\@mkboth\@gobbletwo 160 | \let\@oddhead\@empty 161 | \let\@evenhead\@empty 162 | \let\@evenfoot\copyrightline 163 | \let\@oddfoot\copyrightline} 164 | \def\lipics@opterrshort{Option "\CurrentOption" not supported} 165 | \def\lipics@opterrlong{The option "\CurrentOption" from article.cls is not supported by lipics.cls.} 166 | \DeclareOption{a5paper}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}} 167 | \DeclareOption{b5paper}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}} 168 | \DeclareOption{legalpaper}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}} 169 | \DeclareOption{executivepaper}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}} 170 | \DeclareOption{landscape}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}} 171 | \DeclareOption{10pt}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}} 172 | \DeclareOption{11pt}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}} 173 | \DeclareOption{12pt}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}} 174 | \DeclareOption{oneside}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}} 175 | \DeclareOption{twoside}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}} 176 | \DeclareOption{titlepage}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}} 177 | \DeclareOption{notitlepage}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}} 178 | \DeclareOption{onecolumn}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}} 179 | \DeclareOption{twocolumn}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}} 180 | \DeclareOption{fleqn}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}} 181 | \DeclareOption{openbib}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}} 182 | \DeclareOption{a4paper}{\PassOptionsToClass{\CurrentOption}{article} 183 | \advance\hoffset-2.95mm 184 | \advance\voffset8.8mm} 185 | \DeclareOption{numberwithinsect}{\let\numberwithinsect\relax} 186 | \DeclareOption{cleveref}{\let\usecleveref\relax} 187 | \DeclareOption{autoref}{\let\useautoref\relax} 188 | \DeclareOption{anonymous}{\let\authoranonymous\relax} 189 | \DeclareOption{thm-restate}{\let\usethmrestate\relax} 190 | \DeclareOption{authorcolumns}{\let\authorcolumns\relax} 191 | \DeclareOption*{\PassOptionsToClass{\CurrentOption}{article}} 192 | \ProcessOptions 193 | \LoadClass[twoside,notitlepage,fleqn]{article} 194 | \renewcommand\normalsize{% 195 | \@setfontsize\normalsize\@xpt{13}% 196 | \abovedisplayskip 10\p@ \@plus2\p@ \@minus5\p@ 197 | \abovedisplayshortskip \z@ \@plus3\p@ 198 | \belowdisplayshortskip 6\p@ \@plus3\p@ \@minus3\p@ 199 | \belowdisplayskip \abovedisplayskip 200 | \let\@listi\@listI} 201 | \normalsize 202 | \renewcommand\small{% 203 | \@setfontsize\small\@ixpt{11.5}% 204 | \abovedisplayskip 8.5\p@ \@plus3\p@ \@minus4\p@ 205 | \abovedisplayshortskip \z@ \@plus2\p@ 206 | \belowdisplayshortskip 4\p@ \@plus2\p@ \@minus2\p@ 207 | \def\@listi{\leftmargin\leftmargini 208 | \topsep 4\p@ \@plus2\p@ \@minus2\p@ 209 | \parsep 2\p@ \@plus\p@ \@minus\p@ 210 | \itemsep \parsep}% 211 | \belowdisplayskip \abovedisplayskip 212 | } 213 | \renewcommand\footnotesize{% 214 | \@setfontsize\footnotesize{8.5}{9.5}% 215 | \abovedisplayskip 6\p@ \@plus2\p@ \@minus4\p@ 216 | \abovedisplayshortskip \z@ \@plus\p@ 217 | \belowdisplayshortskip 3\p@ \@plus\p@ \@minus2\p@ 218 | \def\@listi{\leftmargin\leftmargini 219 | \topsep 3\p@ \@plus\p@ \@minus\p@ 220 | \parsep 2\p@ \@plus\p@ \@minus\p@ 221 | \itemsep \parsep}% 222 | \belowdisplayskip \abovedisplayskip 223 | } 224 | \renewcommand\large{\@setfontsize\large{10.5}{13}} 225 | \renewcommand\Large{\@setfontsize\Large{12}{14}} 226 | \setlength\parindent{1.5em} 227 | \setlength\headheight{3mm} 228 | \setlength\headsep {10mm} 229 | \setlength\footskip{3mm} 230 | \setlength\textwidth{140mm} 231 | \setlength\textheight{222mm} 232 | \setlength\oddsidemargin{32mm} 233 | \setlength\evensidemargin{38mm} 234 | \setlength\marginparwidth{25mm} 235 | \setlength\topmargin{13mm} 236 | \setlength{\skip\footins}{2\baselineskip \@plus 4\p@ \@minus 2\p@} 237 | \def\@listi{\leftmargin\leftmargini 238 | \parsep\z@ \@plus\p@ 239 | \topsep 8\p@ \@plus2\p@ \@minus4\p@ 240 | \itemsep \parsep} 241 | \let\@listI\@listi 242 | \@listi 243 | \def\@listii {\leftmargin\leftmarginii 244 | \labelwidth\leftmarginii 245 | \advance\labelwidth-\labelsep 246 | \topsep 4\p@ \@plus2\p@ \@minus\p@ 247 | \parsep\z@ \@plus\p@ 248 | \itemsep \parsep} 249 | \def\@listiii{\leftmargin\leftmarginiii 250 | \labelwidth\leftmarginiii 251 | \advance\labelwidth-\labelsep 252 | \topsep 2\p@ \@plus\p@\@minus\p@ 253 | \parsep \z@ 254 | \partopsep \p@ \@plus\z@ \@minus\p@ 255 | \itemsep \z@ \@plus\p@} 256 | \def\ps@headings{% 257 | \def\@evenhead{\large\sffamily\bfseries 258 | \llap{\hbox to0.5\oddsidemargin{ \ifx\@hideLIPIcs\@undefined\ifx\@ArticleNo\@empty\textcolor{red}{XX}\else\@ArticleNo\fi:\fi\thepage\hss}}\leftmark\hfil}% 259 | \def\@oddhead{\large\sffamily\bfseries\rightmark\hfil 260 | \rlap{\hbox to0.5\oddsidemargin{\hss \ifx\@hideLIPIcs\@undefined\ifx\@ArticleNo\@empty\textcolor{red}{XX}\else\@ArticleNo\fi:\fi\thepage}}}% 261 | \def\@oddfoot{\hfil 262 | \rlap{% 263 | \vtop{% 264 | \vskip10mm 265 | \colorbox{lipicsYellow} 266 | {\@tempdima\evensidemargin 267 | \advance\@tempdima1in 268 | \advance\@tempdima\hoffset 269 | \hb@xt@\@tempdima{% 270 | \ifx\@hideLIPIcs\@undefined 271 | \textcolor{lipicsGray}{\normalsize\sffamily 272 | \bfseries\quad 273 | \expandafter\textsolittle 274 | \expandafter{\@EventShortTitle}}% 275 | \fi 276 | \strut\hss}}}}} 277 | \let\@evenfoot\@empty 278 | \let\@mkboth\markboth 279 | \let\sectionmark\@gobble 280 | \let\subsectionmark\@gobble} 281 | \pagestyle{headings} 282 | \renewcommand\maketitle{\par 283 | \begingroup 284 | \thispagestyle{plain} 285 | \renewcommand\thefootnote{\@fnsymbol\c@footnote}% 286 | \if@twocolumn 287 | \ifnum \col@number=\@ne 288 | \@maketitle 289 | \else 290 | \twocolumn[\@maketitle]% 291 | \fi 292 | \else 293 | \newpage 294 | \global\@topnum\z@ % Prevents figures from going at top of page. 295 | \@maketitle 296 | \fi 297 | \thispagestyle{plain}\@thanks 298 | \endgroup 299 | \global\let\thanks\relax 300 | \global\let\maketitle\relax 301 | \global\let\@maketitle\relax 302 | \global\let\@thanks\@empty 303 | \global\let\@author\@empty 304 | \global\let\@date\@empty 305 | \global\let\@title\@empty 306 | \global\let\title\relax 307 | \global\let\author\relax 308 | \global\let\date\relax 309 | \global\let\and\relax 310 | } 311 | \newwrite\tocfile 312 | \def\@maketitle{% 313 | \newpage 314 | \null\vskip-\baselineskip 315 | \vskip-\headsep 316 | \@titlerunning 317 | \@authorrunning 318 | %%\let \footnote \thanks 319 | \parindent\z@ \raggedright 320 | \if!\@title!\def\@title{\textcolor{red}{Author: Please fill in a title}}\fi 321 | {\LARGE\sffamily\bfseries\mathversion{bold}\@title \par}% 322 | \vskip 1em 323 | \ifx\@author\orig@author 324 | \textcolor{red}{Author: Please provide author information}% 325 | \else 326 | {\def\thefootnote{\@arabic\c@footnote}% 327 | \setcounter{footnote}{0}% 328 | \fontsize{9.5}{12}\selectfont\@author}% 329 | \fi 330 | \bgroup 331 | \immediate\openout\tocfile=\jobname.vtc 332 | \protected@write\tocfile{ 333 | \let\footnote\@gobble 334 | \let\thanks\@gobble 335 | \def\footnotemark{} 336 | \def\and{and }% 337 | \def\,{ } 338 | \def\\{ } 339 | }{% 340 | \string\contitem 341 | \string\title{\@title}% 342 | \string\author{\@authorsfortoc}% 343 | \string\page{\@ArticleNo:\thepage--\@ArticleNo:\number\numexpr\getpagerefnumber{LastPage}}}% 344 | \closeout\tocfile 345 | \egroup 346 | \par} 347 | \renewcommand\tableofcontents{% 348 | \section*{\contentsname}% 349 | \@starttoc{toc}} 350 | \setcounter{secnumdepth}{4} 351 | \renewcommand\section{\@startsection {section}{1}{\z@}% 352 | {-3.5ex \@plus -1ex \@minus -.2ex}% 353 | {2.3ex \@plus.2ex}% 354 | {\sffamily\Large\bfseries\raggedright}} 355 | \renewcommand\subsection{\@startsection{subsection}{2}{\z@}% 356 | {-3.25ex\@plus -1ex \@minus -.2ex}% 357 | {1.5ex \@plus .2ex}% 358 | {\sffamily\Large\bfseries\raggedright}} 359 | \renewcommand\subsubsection{\@startsection{subsubsection}{3}{\z@}% 360 | {-3.25ex\@plus -1ex \@minus -.2ex}% 361 | {1.5ex \@plus .2ex}% 362 | {\sffamily\Large\bfseries\raggedright}} 363 | \renewcommand\paragraph{\@startsection{paragraph}{4}{\z@}% 364 | {-3.25ex \@plus-1ex \@minus-.2ex}% 365 | {1.5ex \@plus .2ex}% 366 | {\sffamily\large\bfseries\raggedright}} 367 | \renewcommand\subparagraph{\@startsection{subparagraph}{5}{\z@}% 368 | {3.25ex \@plus1ex \@minus .2ex}% 369 | {-1em}% 370 | {\sffamily\normalsize\bfseries}} 371 | \setlength\leftmargini \parindent 372 | \setlength\leftmarginii {1.2em} 373 | \setlength\leftmarginiii{1.2em} 374 | \setlength\leftmarginiv {1.2em} 375 | \setlength\leftmarginv {1.2em} 376 | \setlength\leftmarginvi {1.2em} 377 | \renewcommand\labelenumi{% 378 | \textcolor{lipicsGray}{\sffamily\bfseries\upshape\mathversion{bold}\theenumi.}} 379 | \renewcommand\labelenumii{% 380 | \textcolor{lipicsGray}{\sffamily\bfseries\upshape\mathversion{bold}\theenumii.}} 381 | \renewcommand\labelenumiii{% 382 | \textcolor{lipicsGray}{\sffamily\bfseries\upshape\mathversion{bold}\theenumiii.}} 383 | \renewcommand\labelenumiv{% 384 | \textcolor{lipicsGray}{\sffamily\bfseries\upshape\mathversion{bold}\theenumiv.}} 385 | \renewcommand\labelitemi{% 386 | \textcolor{lipicsBulletGray}{\ifnum\@listdepth=\@ne 387 | \rule{0.67em}{0.33em}% 388 | \else 389 | \rule{0.45em}{0.225em}% 390 | \fi}} 391 | \renewcommand\labelitemii{% 392 | \textcolor{lipicsBulletGray}{\rule{0.45em}{0.225em}}} 393 | \renewcommand\labelitemiii{% 394 | \textcolor{lipicsBulletGray}{\sffamily\bfseries\textasteriskcentered}} 395 | \renewcommand\labelitemiv{% 396 | \textcolor{lipicsBulletGray}{\sffamily\bfseries\textperiodcentered}} 397 | \renewenvironment{description} 398 | {\list{}{\advance\partopsep\topsep\topsep\z@\@plus\p@ 399 | \labelwidth\z@ \itemindent-\leftmargin 400 | \let\makelabel\descriptionlabel}} 401 | {\ifnum\@listdepth=\@ne\global\@nobotseplisttrue\fi\endlist} 402 | \renewcommand*\descriptionlabel[1]{% 403 | \hspace\labelsep\textcolor{lipicsGray}{\sffamily\bfseries\mathversion{bold}#1}} 404 | \renewenvironment{abstract}{% 405 | \vskip0.7\bigskipamount 406 | \noindent 407 | \rlap{\color{lipicsLineGray}\vrule\@width\textwidth\@height1\p@}% 408 | \hspace*{7mm}\fboxsep1.5mm\colorbox[rgb]{1,1,1}{\raisebox{-0.4ex}{% 409 | \large\selectfont\sffamily\bfseries\abstractname}}% 410 | \vskip3\p@ 411 | \fontsize{9}{12}\selectfont 412 | \noindent\ignorespaces} 413 | {\vskip0.7\baselineskip\noindent 414 | \subjclassHeading 415 | \ifx\@ccsdescString\@empty 416 | \textcolor{red}{Author: Please fill in 1 or more \string\ccsdesc\space macro}% 417 | \else 418 | \@ccsdescString 419 | \fi 420 | \vskip0.7\baselineskip 421 | \noindent\keywordsHeading 422 | \ifx\@keywords\@empty 423 | \textcolor{red}{Author: Please fill in \string\keywords\space macro}% 424 | \else 425 | \@keywords 426 | \fi 427 | \ifx\@hideLIPIcs\@undefined 428 | \ifx\@DOIPrefix\@empty\else 429 | \vskip0.7\baselineskip\noindent 430 | \doiHeading\href{https://doi.org/\@DOIPrefix.\@EventAcronym.\@EventYear.\@ArticleNo}{\@DOIPrefix.\@EventAcronym.\@EventYear.\@ArticleNo}% 431 | \fi 432 | \fi 433 | \ifx\@category\@empty\else 434 | \vskip0.7\baselineskip\noindent 435 | \categoryHeading\@category 436 | \fi 437 | \ifx\@relatedversion\@empty\else 438 | \vskip0.7\baselineskip\noindent 439 | \relatedversionHeading\@relatedversion 440 | \fi 441 | \ifx\@supplement\@empty\else 442 | \vskip0.7\baselineskip\noindent 443 | \supplementHeading\@supplement 444 | \fi 445 | \ifx\@funding\@empty\else 446 | \vskip0.7\baselineskip\noindent 447 | \fundingHeading\ifx\authoranonymous\relax\textcolor{red}{Anonymous funding}\else\@funding\fi 448 | \fi 449 | \ifx\@acknowledgements\@empty\else 450 | \vskip0.7\baselineskip\noindent 451 | \acknowledgementsHeading\ifx\authoranonymous\relax\textcolor{red}{Anonymous acknowledgements} \else\@acknowledgements\fi 452 | \fi 453 | \protected@write\@auxout{}{\string\gdef\string\@pageNumberEndAbstract{\thepage}}% 454 | }% end abstract 455 | \renewenvironment{thebibliography}[1] 456 | {\if@noskipsec \leavevmode \fi 457 | \par 458 | \@tempskipa-3.5ex \@plus -1ex \@minus -.2ex\relax 459 | \@afterindenttrue 460 | \@tempskipa -\@tempskipa \@afterindentfalse 461 | \if@nobreak 462 | \everypar{}% 463 | \else 464 | \addpenalty\@secpenalty\addvspace\@tempskipa 465 | \fi 466 | \noindent 467 | \rlap{\color{lipicsLineGray}\vrule\@width\textwidth\@height1\p@}% 468 | \hspace*{7mm}\fboxsep1.5mm\colorbox[rgb]{1,1,1}{\raisebox{-0.4ex}{% 469 | \normalsize\sffamily\bfseries\refname}}% 470 | \@xsect{1ex \@plus.2ex}% 471 | \list{\@biblabel{\@arabic\c@enumiv}}% 472 | {\leftmargin8.5mm 473 | \labelsep\leftmargin 474 | \settowidth\labelwidth{\@biblabel{#1}}% 475 | \advance\labelsep-\labelwidth 476 | \usecounter{enumiv}% 477 | \let\p@enumiv\@empty 478 | \renewcommand\theenumiv{\@arabic\c@enumiv}}% 479 | \fontsize{9}{12}\selectfont 480 | \sloppy 481 | \clubpenalty4000 482 | \@clubpenalty \clubpenalty 483 | \widowpenalty4000% 484 | \sfcode`\.\@m\protected@write\@auxout{}{\string\gdef\string\@pageNumberStartBibliography{\thepage}}} 485 | {\def\@noitemerr 486 | {\@latex@warning{Empty `thebibliography' environment}}% 487 | \protected@write\@auxout{}{\string\gdef\string\@pageNumberEndBibliography{\thepage}}% 488 | \endlist} 489 | \g@addto@macro\appendix{\immediate\write\@auxout{\string\gdef\string\@pageNumberStartAppendix{\thepage}}}% 490 | \renewcommand\footnoterule{% 491 | \kern-8\p@ 492 | {\color{lipicsBulletGray}\hrule\@width40mm\@height1\p@}% 493 | \kern6.6\p@} 494 | \renewcommand\@makefntext[1]{% 495 | \parindent\z@\hangindent1em 496 | \leavevmode 497 | \hb@xt@1em{\@makefnmark\hss}#1} 498 | \usepackage{microtype} 499 | \usepackage[utf8]{inputenc} 500 | \IfFileExists{lmodern.sty}{\RequirePackage{lmodern}}{} 501 | \RequirePackage[T1]{fontenc} 502 | \RequirePackage{textcomp} 503 | \RequirePackage[mathscr]{eucal} 504 | \RequirePackage{amssymb} 505 | \PassOptionsToPackage{retainmissing}{MnSymbol} 506 | \AtBeginDocument{\@ifpackageloaded{MnSymbol}% 507 | {\expandafter\let\csname ver@amssymb.sty\endcsname\relax 508 | \let\complement\@undefined 509 | \RequirePackage{amssymb}}{}} 510 | \RequirePackage{soul} 511 | \sodef\textsolittle{}{.12em}{.5em\@plus.08em\@minus.06em}% 512 | {.4em\@plus.275em\@minus.183em} 513 | \RequirePackage{color} %kept for backward compatibility 514 | \AtBeginDocument{ 515 | \@ifpackageloaded{xcolor}{ 516 | }{ 517 | \RequirePackage{xcolor} 518 | } 519 | \definecolor{darkgray}{rgb}{0.31,0.31,0.33} 520 | \definecolor[named]{lipicsGray}{rgb}{0.31,0.31,0.33} 521 | \definecolor[named]{lipicsBulletGray}{rgb}{0.60,0.60,0.61} 522 | \definecolor[named]{lipicsLineGray}{rgb}{0.51,0.50,0.52} 523 | \definecolor[named]{lipicsLightGray}{rgb}{0.85,0.85,0.86} 524 | \definecolor[named]{lipicsYellow}{rgb}{0.99,0.78,0.07} 525 | } 526 | \RequirePackage{babel} 527 | \RequirePackage[tbtags,fleqn]{amsmath} 528 | \AtBeginDocument{ 529 | \@ifpackageloaded{enumitem}{\ClassWarning{Package 'enuitem' incompatible}{Don't use package 'enumitem'; Package enumerate preloaded!}}{} 530 | \@ifpackageloaded{paralist}{\ClassWarning{Package 'paralist' incompatible}{Don't use package 'paralist'; Package enumerate preloaded!}}{} 531 | } 532 | \RequirePackage{enumerate} 533 | \def\@enum@{\list{\textcolor{lipicsGray}{\sffamily\bfseries\upshape\mathversion{bold}\csname label\@enumctr\endcsname}}% 534 | {\advance\partopsep\topsep 535 | \topsep\z@\@plus\p@ 536 | \usecounter{\@enumctr}\def\makelabel##1{\hss\llap{##1}}}} 537 | \def\romanenumerate{\enumerate[(i)]} 538 | \let\endromanenumerate\endenumerate 539 | \def\alphaenumerate{\enumerate[(a)]} 540 | \let\endalphaenumerate\endenumerate 541 | \def\bracketenumerate{\enumerate[(1)]} 542 | \let\endbracketenumerate\endenumerate 543 | \RequirePackage{graphicx} 544 | \RequirePackage{array} 545 | \let\@classzold\@classz 546 | \def\@classz{% 547 | \expandafter\ifx\d@llarbegin\begingroup 548 | \toks \count@ = 549 | \expandafter{\expandafter\small\the\toks\count@}% 550 | \fi 551 | \@classzold} 552 | \RequirePackage{multirow} 553 | \RequirePackage{tabularx} 554 | \RequirePackage[online]{threeparttable} 555 | \def\TPTtagStyle#1{#1)} 556 | \def\tablenotes{\small\TPT@defaults 557 | \@ifnextchar[\TPT@setuptnotes\TPTdoTablenotes} % ] 558 | \RequirePackage{listings} 559 | \lstset{basicstyle=\small\ttfamily,% 560 | backgroundcolor=\color{lipicsLightGray},% 561 | frame=single,framerule=0pt,xleftmargin=\fboxsep,xrightmargin=\fboxsep} 562 | \RequirePackage[left,mathlines]{lineno} 563 | \linenumbers 564 | \renewcommand\linenumberfont{\normalfont\tiny\sffamily} 565 | %%%% patch to cope with amsmath 566 | %%%% http://phaseportrait.blogspot.de/2007/08/lineno-and-amsmath-compatibility.html 567 | \newcommand*\patchAmsMathEnvironmentForLineno[1]{% 568 | \expandafter\let\csname old#1\expandafter\endcsname\csname #1\endcsname 569 | \expandafter\let\csname oldend#1\expandafter\endcsname\csname end#1\endcsname 570 | \renewenvironment{#1}% 571 | {\linenomath\csname old#1\endcsname}% 572 | {\csname oldend#1\endcsname\endlinenomath}}% 573 | \newcommand*\patchBothAmsMathEnvironmentsForLineno[1]{% 574 | \patchAmsMathEnvironmentForLineno{#1}% 575 | \patchAmsMathEnvironmentForLineno{#1*}}% 576 | \AtBeginDocument{% 577 | \patchBothAmsMathEnvironmentsForLineno{equation}% 578 | \patchBothAmsMathEnvironmentsForLineno{align}% 579 | \patchBothAmsMathEnvironmentsForLineno{flalign}% 580 | \patchBothAmsMathEnvironmentsForLineno{alignat}% 581 | \patchBothAmsMathEnvironmentsForLineno{gather}% 582 | \patchBothAmsMathEnvironmentsForLineno{multline}} 583 | \RequirePackage{lastpage} 584 | \RequirePackage{hyperref} 585 | \AtBeginDocument{ 586 | \pdfstringdefDisableCommands{% 587 | \let\thanks\@gobble 588 | \let\footnote\@gobble 589 | \def\footnotemark{} 590 | \def\cs#1{\textbackslash #1}% 591 | \let\normalfont\@empty 592 | \let\scshape\@empty 593 | \def\and{and }% 594 | \def\,{ } 595 | } 596 | \hypersetup{ 597 | breaklinks=true, 598 | pdfborder={0 0 0}, 599 | pdftitle={\@title}, 600 | pdfauthor={\ifx\authoranonymous\relax Anonymous author(s) \else \@authorsforpdf \fi}, 601 | pdfkeywords={\@keywords}, 602 | pdfcreator={LaTeX with lipics-v2019.cls and hyperref.sty}, 603 | %TODO Find way to move information to correct fields in pdf metadata 604 | pdfsubject={LIPIcs, Vol.\@SeriesVolume, \@EventShortTitle; Copyright (C) \ifx\authoranonymous\relax Anonymous author(s) \else \@copyrightholder; \fi licensed under Creative Commons License CC-BY} 605 | %pdflicenseurl={https://creativecommons.org/licenses/by/3.0/}, 606 | %pdfcopyright={Copyright (C), \@copyrightholder; licensed under Creative Commons License CC-BY}, 607 | } 608 | } 609 | \RequirePackage[labelsep=space,singlelinecheck=false,% 610 | font={up,small},labelfont={sf,bf},% 611 | listof=false]{caption}%"listof" instead of "list" for backward compatibility 612 | \@ifpackagelater{hyperref}{2009/12/09} 613 | {\captionsetup{compatibility=false}}%cf. http://groups.google.de/group/comp.text.tex/browse_thread/thread/db9310eb540fbbd8/42e30f3b7b3aa17a?lnk=raot 614 | {} 615 | \DeclareCaptionLabelFormat{boxed}{% 616 | \kern0.05em{\color[rgb]{0.99,0.78,0.07}\rule{0.73em}{0.73em}}% 617 | \hspace*{0.67em}\bothIfFirst{#1}{~}#2} 618 | \captionsetup{labelformat=boxed} 619 | \captionsetup[table]{position=top} 620 | \RequirePackage[figuresright]{rotating} 621 | \caption@AtBeginDocument{\@ifpackageloaded{subfig}{\ClassError{lipics}{% 622 | Do not load the subfig package}{The more recent subcaption package is already loaded}}{}} 623 | \RequirePackage{subcaption} 624 | \def\titlerunning#1{\gdef\@titlerunning{{\let\footnote\@gobble\markboth{#1}{#1}}}} 625 | \def\authorrunning#1{% 626 | \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}}} 627 | \titlerunning{\@title} 628 | \authorrunning{\textcolor{red}{Author: Please use the \string\authorrunning\space macro}} 629 | \def\EventLongTitle#1{\gdef\@EventLongTitle{#1}} 630 | \EventLongTitle{} 631 | \def\EventShortTitle#1{\gdef\@EventShortTitle{#1}} 632 | \EventShortTitle{} 633 | \def\EventEditors#1{\gdef\@EventEditors{#1}} 634 | \EventEditors{} 635 | \def\EventNoEds#1{\gdef\@EventNoEds{#1}\xdef\@Eds{Editor\ifnum#1>1s\fi}} 636 | \EventNoEds{1} 637 | \def\EventLogo#1{\gdef\@EventLogo{#1}} 638 | \EventLogo{} 639 | \def\EventAcronym#1{\gdef\@EventAcronym{#1}} 640 | \EventAcronym{} 641 | \def\EventYear#1{\gdef\@EventYear{#1}} 642 | \EventYear{} 643 | \def\EventDate#1{\gdef\@EventDate{#1}} 644 | \EventDate{} 645 | \def\EventLocation#1{\gdef\@EventLocation{#1}} 646 | \EventLocation{} 647 | \def\SeriesVolume#1{\gdef\@SeriesVolume{#1}} 648 | \SeriesVolume{} 649 | \def\ArticleNo#1{\gdef\@ArticleNo{#1}} 650 | \ArticleNo{} 651 | \def\DOIPrefix#1{\gdef\@DOIPrefix{#1}} 652 | \DOIPrefix{10.4230/LIPIcs} 653 | \def\and{\newline} 654 | \let\orig@author\@author 655 | \let\@authorsfortoc\@empty 656 | \let\@authorsforpdf\@empty 657 | \newcount\c@author 658 | \newcounter{currentauthor} 659 | \def\@authornum{0} 660 | \def\author#1#2#3#4#5{% 661 | \ifx\@author\orig@author\let\@author\@empty\fi 662 | \g@addto@macro\@author{% 663 | \ifx\authorcolumns\relax 664 | \ifnum\c@author>6 665 | \stepcounter{currentauthor} 666 | \ifodd\value{currentauthor} 667 | \begin{minipage}[t]{\textwidth} 668 | \begin{minipage}[t]{0.49\textwidth} 669 | \else 670 | \hfill \begin{minipage}[t]{0.49\textwidth} 671 | \fi 672 | \else 673 | \ClassWarning{Option 'authorcolumns' only applicable for > 6 authors}{Option 'authorcolumns' only applicable for >6 authors!} 674 | \addvspace{0.5\baselineskip} 675 | \fi 676 | \else 677 | \addvspace{0.5\baselineskip} 678 | \fi 679 | {\Large\bfseries 680 | \if!#1! 681 | \textcolor{red}{Author: Please enter author name}% 682 | \else 683 | \ifx\authoranonymous\relax 684 | \textcolor{red}{Anonymous author} 685 | \else 686 | #1% 687 | \if!#4!\else{\space\href{#4}{\includegraphics[height=9\p@]{orcid}}}\fi 688 | \if!#5!\else 689 | \ifx\@funding\@empty 690 | \expandafter\g@addto@macro\expandafter\@funding{\textit{\expandafter{\let\footnote\@gobble #1}}:\space{#5}} 691 | \else 692 | \expandafter\g@addto@macro\expandafter\@funding{\\\textit{\expandafter{\let\footnote\@gobble #1}}:\space{#5}} 693 | \fi 694 | \fi 695 | \fi 696 | \fi 697 | } 698 | {\small 699 | \if!#2!\textcolor{red}{Author: Please enter affiliation as second parameter of the author macro}\else{\\* \ifx\authoranonymous\relax\textcolor{red}{Anonymous affiliation}\else#2\fi}\fi 700 | \if!#3!\else{\ifx\authoranonymous\relax\else\\*\href{mailto:#3}{#3}\fi}\fi 701 | }\par 702 | \ifx\authorcolumns\relax 703 | \ifnum\c@author>6 704 | \end{minipage} 705 | \ifnum\c@author=\value{currentauthor} 706 | \end{minipage} 707 | \else 708 | \ifodd\value{currentauthor} 709 | \else 710 | \end{minipage}% 711 | \medskip 712 | \fi 713 | \fi 714 | \fi 715 | \fi}% 716 | \global\advance\c@author\@ne 717 | \protected@write\@auxout{}{\string\gdef\string\@authornum{\the\c@author}} 718 | \ifnum\c@author=\@ne 719 | \gdef\@authorsfortoc{#1}% 720 | \gdef\@authorsforpdf{#1} 721 | \else 722 | \expandafter\g@addto@macro\expandafter\@authorsforpdf\expandafter{,\space #1} 723 | \expandafter\g@addto@macro\expandafter\@authorsfortoc\expandafter{\expandafter\csname\the\c@author authand\endcsname#1}% 724 | \@namedef{\the\c@author authand}{,\space}% 725 | \AtBeginDocument{% 726 | \expandafter\ifnum\@authornum=2 727 | \@namedef{2authand}{\space and\space}% 728 | \else 729 | \@namedef{\@authornum authand}{,\space and\space}% 730 | \fi} 731 | \fi} 732 | \newcommand*\affil[2][]{% 733 | \ClassError{lipics} 734 | {\string\affil\space deprecated: Please enter affiliation as second parameter of the author macro} 735 | {Since 2017, \string\affil\space is obsolete in lipics.}} 736 | \newcommand*\Copyright[1]{% 737 | \def\@copyrightholder{#1} 738 | \def\@Copyright{% 739 | \setbox\@tempboxa\hbox{\includegraphics[height=14\p@,clip]{cc-by}}% 740 | \@rightskip\@flushglue \rightskip\@rightskip 741 | \hangindent\dimexpr\wd\@tempboxa+0.5em\relax 742 | \href{https://creativecommons.org/licenses/by/3.0/}% 743 | {\smash{\lower\baselineskip\hbox{\unhcopy\@tempboxa}}}\enskip 744 | \textcopyright\ % 745 | \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 746 | ;\\% 747 | licensed under Creative Commons License CC-BY\ifx!#1!\\\null\fi\par}} 748 | \Copyright{\textcolor{red}{Author: Please provide a copyright holder}} 749 | \let\@copyrightholder\@empty 750 | \def\hideLIPIcs{\let\@hideLIPIcs\relax} 751 | \def\keywords#1{\def\@keywords{#1}} 752 | \let\@keywords\@empty 753 | \def\keywordsHeading{% 754 | \textcolor{lipicsGray}{\fontsize{9}{12}\sffamily\bfseries 755 | Keywords and phrases\enskip}} 756 | \RequirePackage{comment} 757 | \excludecomment{CCSXML} 758 | % inspired by https://tex.stackexchange.com/questions/12810/how-do-i-split-a-string 759 | \global\newcommand\ccsdesc[2][100]{\@ccsdesc#1~#2~~\relax} 760 | \usepackage{xstring} 761 | \let\orig@ccsdesc\@ccsdesc 762 | \let\@ccsdesc\@empty 763 | \let\@ccsdescString\@empty 764 | \gdef\@ccsdesc#1~#2~#3~{ 765 | \ifx\@ccsdesc\orig@ccsdesc\let\@ccsdesc\@empty\fi 766 | \ifx!#3! 767 | \ifx\@ccsdescString\@empty 768 | \g@addto@macro\@ccsdescString{{#2}} 769 | \else 770 | \g@addto@macro\@ccsdescString{; {#2}} 771 | \fi 772 | \else 773 | \ifx\@ccsdescString\@empty 774 | \g@addto@macro\@ccsdescString{{#2} $\rightarrow$ {#3}} 775 | \else 776 | \g@addto@macro\@ccsdescString{; {#2} $\rightarrow$ {#3}} 777 | \fi 778 | \fi 779 | \ccsdescEnd 780 | } 781 | \def\ccsdescEnd#1\relax{} 782 | \def\subjclass#1{ 783 | \ClassError{lipics} 784 | {\string\subjclass\space deprecated: Please enter subject classification in 1 or more ccsdesc macros} 785 | {Since 2019, \string\subjclass\space is obsolete in lipics.}} 786 | \let\@subjclass\@empty 787 | \def\subjclassHeading{% 788 | \textcolor{lipicsGray}{\fontsize{9}{12}\sffamily\bfseries 789 | 2012 ACM Subject Classification\enskip}} 790 | \def\doiHeading{% 791 | \textcolor{lipicsGray}{\fontsize{9}{12}\sffamily\bfseries 792 | Digital Object Identifier\enskip}} 793 | \def\category#1{\def\@category{#1}} 794 | \let\@category\@empty 795 | \def\categoryHeading{% 796 | \textcolor{lipicsGray}{\fontsize{9}{12}\sffamily\bfseries 797 | Category\enskip}} 798 | \def\relatedversion#1{\def\@relatedversion{#1}} 799 | \let\@relatedversion\@empty 800 | \def\relatedversionHeading{% 801 | \textcolor{lipicsGray}{\fontsize{9}{12}\sffamily\bfseries 802 | Related Version\enskip}} 803 | \def\supplement#1{\def\@supplement{#1}} 804 | \let\@supplement\@empty 805 | \def\supplementHeading{% 806 | \textcolor{lipicsGray}{\fontsize{9}{12}\sffamily\bfseries 807 | Supplementary Material\enskip}} 808 | \def\funding#1{\def\@funding{#1}} 809 | \let\@funding\@empty 810 | \def\fundingHeading{% 811 | \textcolor{lipicsGray}{\fontsize{9}{12}\sffamily\bfseries 812 | Funding\enskip}} 813 | \def\acknowledgements#1{\def\@acknowledgements{#1}} 814 | \let\@acknowledgements\@empty 815 | \def\acknowledgementsHeading{% 816 | \textcolor{lipicsGray}{\fontsize{9}{12}\sffamily\bfseries 817 | Acknowledgements\enskip}} 818 | \RequirePackage{amsthm} 819 | \ifx\usethmrestate\relax 820 | \RequirePackage{thm-restate} 821 | \fi 822 | \thm@headfont{% 823 | \textcolor{lipicsGray}{$\blacktriangleright$}\nobreakspace\sffamily\bfseries} 824 | \def\th@remark{% 825 | \thm@headfont{% 826 | \textcolor{lipicsGray}{$\blacktriangleright$}\nobreakspace\sffamily}% 827 | \normalfont % body font 828 | \thm@preskip\topsep \divide\thm@preskip\tw@ 829 | \thm@postskip\thm@preskip 830 | } 831 | \def\@endtheorem{\endtrivlist}%\@endpefalse 832 | \renewcommand\qedsymbol{\textcolor{lipicsGray}{\ensuremath{\blacktriangleleft}}} 833 | \renewenvironment{proof}[1][\proofname]{\par 834 | \pushQED{\qed}% 835 | \normalfont \topsep6\p@\@plus6\p@\relax 836 | \trivlist 837 | \item[\hskip\labelsep 838 | \color{lipicsGray}\sffamily\bfseries 839 | #1\@addpunct{.}]\ignorespaces 840 | }{% 841 | \popQED\endtrivlist%\@endpefalse 842 | } 843 | \newenvironment{claimproof}[1][\proofname]{ 844 | \pushQED{\qed}% 845 | \normalfont \topsep6\p@\@plus6\p@\relax 846 | \trivlist 847 | \item[\hskip\labelsep 848 | \color{lipicsGray}\sffamily 849 | #1\@addpunct{.}]\ignorespaces 850 | }{% 851 | \renewcommand\qedsymbol{\textcolor{lipicsGray}{\ensuremath{\vartriangleleft}}} 852 | \popQED\endtrivlist%\@endpefalse 853 | \renewcommand\qedsymbol{\textcolor{lipicsGray}{\ensuremath{\blacktriangleleft}}} 854 | } 855 | % inspired by qed of amsthm class 856 | \DeclareRobustCommand{\lipicsEnd}{% 857 | \leavevmode\unskip\penalty9999 \hbox{}\nobreak\hfill 858 | \quad\hbox{$\lrcorner$}% 859 | } 860 | \AtBeginDocument{ 861 | \@ifpackageloaded{algorithm2e}{ 862 | \renewcommand{\algorithmcfname}{\sffamily\bfseries{}Algorithm} 863 | \SetAlgoCaptionSeparator{~} 864 | \SetAlCapHSkip{0pt} 865 | \renewcommand{\algocf@captiontext}[2]{% 866 | \kern0.05em{\color{lipicsYellow}\rule{0.73em}{0.73em}}% 867 | \hspace*{0.67em}\small #1\algocf@capseparator\nobreakspace#2} 868 | \renewcommand{\algocf@makecaption}[2]{% 869 | \parbox[t]{\textwidth}{\algocf@captiontext{#1}{#2}}% 870 | }% 871 | \renewcommand{\@algocf@capt@boxed}{above} 872 | \renewcommand{\@algocf@capt@ruled}{above} 873 | \setlength\algotitleheightrule{0pt} 874 | }{} 875 | \@ifpackageloaded{algorithm}{ 876 | \captionsetup[algorithm]{name=Algorithm, labelformat=boxed, position=top} 877 | \newcommand\fs@ruled@notop{\def\@fs@cfont{\bfseries}\let\@fs@capt\floatc@ruled 878 | \def\@fs@pre{}% 879 | \def\@fs@post{\kern2pt\hrule\relax}% 880 | \def\@fs@mid{\kern2pt\hrule\kern2pt}% 881 | \let\@fs@iftopcapt\iftrue} 882 | \renewcommand\fst@algorithm{\fs@ruled@notop} 883 | }{} 884 | \ifx\usecleveref\relax\else 885 | \@ifpackageloaded{cleveref}{\ClassWarning{Use document option 'cleveref' instead}{Use document option 'cleveref' instead directly loading package 'cleveref'}}{} 886 | \fi 887 | \ifx\usethmrestate\relax\else 888 | \@ifpackageloaded{thm-restate}{\ClassWarning{Use document option 'thm-restate' instead}{Use document option 'thm-restate' instead directly loading package 'thm-restate'}}{} 889 | \fi 890 | \ifx\useautoref\relax 891 | \@ifundefined{algorithmautorefname}{\newcommand{\algorithmautorefname}{Algorithm}}{\renewcommand{\algorithmautorefname}{Algorithm}}% 892 | \fi 893 | } 894 | 895 | \ifx\usecleveref\relax 896 | \RequirePackage[capitalise, noabbrev]{cleveref} 897 | \crefname{algocf}{Algorithm}{Algorithms} 898 | \Crefname{algocf}{Algorithm}{Algorithms} 899 | \fi 900 | \ifx\useautoref\relax 901 | \RequirePackage{aliascnt} 902 | \fi 903 | \newtheoremstyle{claimstyle}{\topsep}{\topsep}{}{0pt}{\sffamily}{. }{5pt plus 1pt minus 1pt}% 904 | {$\vartriangleright$ \thmname{#1}\thmnumber{ #2}\thmnote{ (#3)}} 905 | \theoremstyle{plain} 906 | \newtheorem{theorem}{Theorem} 907 | \ifx\numberwithinsect\relax 908 | \@addtoreset{theorem}{section} 909 | \expandafter\def\expandafter\thetheorem\expandafter{% 910 | \expandafter\thesection\expandafter\@thmcountersep\thetheorem} 911 | \fi 912 | 913 | \ifx\useautoref\relax 914 | \addto\extrasenglish{% 915 | \def\chapterautorefname{Chapter}% 916 | \def\sectionautorefname{Section}% 917 | \def\subsectionautorefname{Subsection}% 918 | \def\subsubsectionautorefname{Subsubsection}% 919 | \def\paragraphautorefname{Paragraph}% 920 | \def\subparagraphautorefname{Subparagraph}% 921 | } 922 | \addto\extrasUKenglish{% 923 | \def\chapterautorefname{Chapter}% 924 | \def\sectionautorefname{Section}% 925 | \def\subsectionautorefname{Subsection}% 926 | \def\subsubsectionautorefname{Subsubsection}% 927 | \def\paragraphautorefname{Paragraph}% 928 | \def\subparagraphautorefname{Subparagraph}% 929 | } 930 | \addto\extrasUSenglish{% 931 | \def\chapterautorefname{Chapter}% 932 | \def\sectionautorefname{Section}% 933 | \def\subsectionautorefname{Subsection}% 934 | \def\subsubsectionautorefname{Subsubsection}% 935 | \def\paragraphautorefname{Paragraph}% 936 | \def\subparagraphautorefname{Subparagraph}% 937 | } 938 | \ifx\usethmrestate\relax 939 | \newtheorem{lemma}[theorem]{Lemma} 940 | \newtheorem{corollary}[theorem]{Corollary} 941 | \newtheorem{proposition}[theorem]{Proposition} 942 | \newtheorem{exercise}[theorem]{Exercise} 943 | \newtheorem{definition}[theorem]{Definition} 944 | \theoremstyle{definition} 945 | \newtheorem{example}[theorem]{Example} 946 | \theoremstyle{remark} 947 | \newtheorem{note}[theorem]{Note} 948 | \newtheorem*{note*}{Note} 949 | \newtheorem{remark}[theorem]{Remark} 950 | \newtheorem*{remark*}{Remark} 951 | \theoremstyle{claimstyle} 952 | \newtheorem{claim}[theorem]{Claim} 953 | \newtheorem*{claim*}{Claim} 954 | \else 955 | \newaliascnt{lemma}{theorem} 956 | \newtheorem{lemma}[lemma]{Lemma} 957 | \aliascntresetthe{lemma} 958 | \newcommand{\lemmaautorefname}{Lemma} 959 | \newaliascnt{corollary}{theorem} 960 | \newtheorem{corollary}[corollary]{Corollary} 961 | \aliascntresetthe{corollary} 962 | \newcommand{\corollaryautorefname}{Corollary} 963 | \newaliascnt{proposition}{theorem} 964 | \newtheorem{proposition}[proposition]{Proposition} 965 | \aliascntresetthe{proposition} 966 | \newcommand{\propositionautorefname}{Proposition} 967 | \newaliascnt{exercise}{theorem} 968 | \newtheorem{exercise}[exercise]{Exercise} 969 | \aliascntresetthe{exercise} 970 | \newcommand{\exerciseautorefname}{Exercise} 971 | \newaliascnt{definition}{theorem} 972 | \newtheorem{definition}[definition]{Definition} 973 | \aliascntresetthe{definition} 974 | \newcommand{\definitionautorefname}{Definition} 975 | \theoremstyle{definition} 976 | \newaliascnt{example}{theorem} 977 | \newtheorem{example}[example]{Example} 978 | \aliascntresetthe{example} 979 | \newcommand{\exampleautorefname}{Example} 980 | \theoremstyle{remark} 981 | \newaliascnt{note}{theorem} 982 | \newtheorem{note}[note]{Note} 983 | \aliascntresetthe{note} 984 | \newcommand{\noteautorefname}{Note} 985 | \newtheorem*{note*}{Note} 986 | \newaliascnt{remark}{theorem} 987 | \newtheorem{remark}[remark]{Remark} 988 | \aliascntresetthe{remark} 989 | \newcommand{\remarkautorefname}{Remark} 990 | \newtheorem*{remark*}{Remark} 991 | \theoremstyle{claimstyle} 992 | \newaliascnt{claim}{theorem} 993 | \newtheorem{claim}[claim]{Claim} 994 | \aliascntresetthe{claim} 995 | \newcommand{\claimautorefname}{Claim} 996 | \newtheorem*{claim*}{Claim} 997 | \fi 998 | \else 999 | \newtheorem{lemma}[theorem]{Lemma} 1000 | \newtheorem{corollary}[theorem]{Corollary} 1001 | \newtheorem{proposition}[theorem]{Proposition} 1002 | \newtheorem{exercise}[theorem]{Exercise} 1003 | \newtheorem{definition}[theorem]{Definition} 1004 | \theoremstyle{definition} 1005 | \newtheorem{example}[theorem]{Example} 1006 | \theoremstyle{remark} 1007 | \newtheorem{note}[theorem]{Note} 1008 | \newtheorem*{note*}{Note} 1009 | \newtheorem{remark}[theorem]{Remark} 1010 | \newtheorem*{remark*}{Remark} 1011 | \theoremstyle{claimstyle} 1012 | \newtheorem{claim}[theorem]{Claim} 1013 | \newtheorem*{claim*}{Claim} 1014 | \fi 1015 | \theoremstyle{plain} 1016 | \endinput 1017 | %% 1018 | %% End of file `lipics-v2019.cls'. 1019 | -------------------------------------------------------------------------------- /paper/orcid.pdf: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/jashug/WhyNotW/c1f26bde6f45519253ed766b99136649eaf9b5bb/paper/orcid.pdf -------------------------------------------------------------------------------- /paper/whynotw.bib: -------------------------------------------------------------------------------- 1 | @phdthesis{hoffman-thesis, 2 | author = {Martin Hoffman}, 3 | title = {Extensional Constructs in Intensional Type Theory}, 4 | year = {1995}, 5 | school = {University of Edinburgh}, 6 | } 7 | 8 | @inproceedings{altenkirch-exteq, 9 | author = {Thorsten Altenkirch}, 10 | title = {Extensional Equality in Intensional Type Theory}, 11 | booktitle = {14th Annual {IEEE} Symposium on Logic in Computer Science, Trento, 12 | Italy, July 2-5, 1999}, 13 | pages = {412--420}, 14 | publisher = {{IEEE} Computer Society}, 15 | year = {1999}, 16 | url = {https://doi.org/10.1109/LICS.1999.782636}, 17 | doi = {10.1109/LICS.1999.782636}, 18 | timestamp = {Wed, 16 Oct 2019 14:14:54 +0200}, 19 | biburl = {https://dblp.org/rec/conf/lics/Altenkirch99.bib}, 20 | bibsource = {dblp computer science bibliography, https://dblp.org} 21 | } 22 | 23 | @article{dybjer-generic, 24 | author = {Marcin Benke and 25 | Peter Dybjer and 26 | Patrik Jansson}, 27 | title = {Universes for Generic Programs and Proofs in Dependent Type Theory}, 28 | journal = {Nord. J. Comput.}, 29 | volume = {10}, 30 | number = {4}, 31 | pages = {265--289}, 32 | year = {2003}, 33 | timestamp = {Wed, 14 Jan 2004 15:34:07 +0100}, 34 | biburl = {https://dblp.org/rec/journals/njc/BenkeDJ03.bib}, 35 | bibsource = {dblp computer science bibliography, https://dblp.org} 36 | } 37 | 38 | @misc{palmgren2019setoids, 39 | title={From type theory to setoids and back}, 40 | author={Erik Palmgren}, 41 | year={2019}, 42 | eprint={1909.01414}, 43 | archivePrefix={arXiv}, 44 | primaryClass={math.LO} 45 | } 46 | 47 | @article{dybjer-W-encodes-inductive, 48 | title = "Representing inductively defined sets by wellorderings in {M}artin-{L}öf's type theory", 49 | journal = "Theoretical Computer Science", 50 | volume = "176", 51 | number = "1", 52 | pages = "329 - 335", 53 | year = "1997", 54 | issn = "0304-3975", 55 | doi = "https://doi.org/10.1016/S0304-3975(96)00145-4", 56 | author = "Peter Dybjer", 57 | abstract = "We prove that every strictly positive endofunctor on the category of sets generated by Martin-Löf's extensional type theory has an initial algebra. This representation of inductively defined sets uses essentially the wellorderings introduced by Martin-Löf in “Constructive Mathematics and Computer Programming”.", 58 | annotation = "Construction of inductive types from W types in extensional type theory.", 59 | } 60 | 61 | @misc{mcbride_wtypes, 62 | title={W-types: good news and bad news}, 63 | url={https://mazzo.li/epilogue/index.html%3Fp=324.html}, 64 | journal={Epilogue for Epigram}, 65 | author={McBride, Conor}, 66 | year={2010}, 67 | month={Mar}, 68 | annotation = "Conor McBride on the unsuitability of W types for inductive data.", 69 | } 70 | 71 | @article{luo-wellordering, 72 | author = {Goguen, Healfdene and Luo, Zhaohui}, 73 | year = {1992}, 74 | pages = {}, 75 | title = {Inductive Data Types: Well-ordering Types Revisited}, 76 | journal = {Logical Environments}, 77 | url={https://www.cs.rhul.ac.uk/home/zhaohui/WTYPES93.pdf}, 78 | } 79 | 80 | @book{programming-in-MLTT, 81 | author = {Nordsrtöm, Bengt and Petersson, Kent and Smith, Jan}, 82 | year = {1990}, 83 | pages = {}, 84 | title = {Programming in Martin-Löf's Type Theory}, 85 | publisher = {Oxford University Press}, 86 | } 87 | 88 | @Book{mltt, 89 | title = {Intuitionistic Type Theory}, 90 | author = {Per Martin-L\"of}, 91 | publisher = {Bibliopolis}, 92 | year = {1984}, 93 | note = {Notes by G. Sambin of a series of lectures given in Padua, 1980}, 94 | annotation = "Generic reference for dependent type theory.", 95 | } 96 | 97 | @incollection{cpdt-moredep, 98 | author = {}, 99 | isbn = {9780262317870}, 100 | title = "{More Dependent Types}", 101 | booktitle = "{Certified Programming with Dependent Types: A Pragmatic Introduction to the Coq Proof Assistant}", 102 | publisher = {The MIT Press}, 103 | year = {2013}, 104 | month = {12}, 105 | doi = {10.7551/mitpress/9153.003.0011}, 106 | url = {https://doi.org/10.7551/mitpress/9153.003.0011}, 107 | eprint = {https://direct.mit.edu/book/chapter-pdf/268879/9780262317870\_cah.pdf}, 108 | annotation = "Describes the convoy pattern.", 109 | } 110 | 111 | @inproceedings{levitation, 112 | author = {James Chapman and 113 | Pierre{-}{\'{E}}variste Dagand and 114 | Conor McBride and 115 | Peter Morris}, 116 | editor = {Paul Hudak and 117 | Stephanie Weirich}, 118 | title = {The gentle art of levitation}, 119 | booktitle = {Proceeding of the 15th {ACM} {SIGPLAN} international conference on 120 | Functional programming, {ICFP} 2010, Baltimore, Maryland, USA, September 121 | 27-29, 2010}, 122 | pages = {3--14}, 123 | publisher = {{ACM}}, 124 | year = {2010}, 125 | url = {https://doi.org/10.1145/1863543.1863547}, 126 | doi = {10.1145/1863543.1863547}, 127 | timestamp = {Tue, 06 Nov 2018 16:59:24 +0100}, 128 | biburl = {https://dblp.org/rec/conf/icfp/ChapmanDMM10.bib}, 129 | bibsource = {dblp computer science bibliography, https://dblp.org}, 130 | annotation = "Levitated universe of inductive types. Preprint url " 131 | } 132 | 133 | @inproceedings{finite-axiom-IR, 134 | author = {Peter Dybjer and 135 | Anton Setzer}, 136 | editor = {Jean{-}Yves Girard}, 137 | title = {A Finite Axiomatization of Inductive-Recursive Definitions}, 138 | booktitle = {Typed Lambda Calculi and Applications, 4th International Conference, 139 | TLCA'99, L'Aquila, Italy, April 7-9, 1999, Proceedings}, 140 | series = {Lecture Notes in Computer Science}, 141 | volume = {1581}, 142 | pages = {129--146}, 143 | publisher = {Springer}, 144 | year = {1999}, 145 | url = {https://doi.org/10.1007/3-540-48959-2\_11}, 146 | doi = {10.1007/3-540-48959-2\_11}, 147 | timestamp = {Tue, 14 May 2019 10:00:41 +0200}, 148 | biburl = {https://dblp.org/rec/conf/tlca/DybjerS99.bib}, 149 | bibsource = {dblp computer science bibliography, https://dblp.org}, 150 | annotation = "Codes for IR types, non-greek constructor names." 151 | } 152 | 153 | @InProceedings{cchm-cubical, 154 | author = {Cyril Cohen and Thierry Coquand and Simon Huber and Anders M{\"o}rtberg}, 155 | title = {{Cubical Type Theory: A Constructive Interpretation of the Univalence Axiom}}, 156 | booktitle = {21st International Conference on Types for Proofs and Programs (TYPES 2015)}, 157 | pages = {5:1--5:34}, 158 | series = {Leibniz International Proceedings in Informatics (LIPIcs)}, 159 | ISBN = {978-3-95977-030-9}, 160 | ISSN = {1868-8969}, 161 | year = {2018}, 162 | volume = {69}, 163 | editor = {Tarmo Uustalu}, 164 | publisher = {Schloss Dagstuhl--Leibniz-Zentrum fuer Informatik}, 165 | address = {Dagstuhl, Germany}, 166 | URL = {http://drops.dagstuhl.de/opus/volltexte/2018/8475}, 167 | URN = {urn:nbn:de:0030-drops-84754}, 168 | doi = {10.4230/LIPIcs.TYPES.2015.5}, 169 | annotation = "First presentation of syntactic cubical type theory. Includes connections, reversals, Glue types.", 170 | } 171 | 172 | @inproceedings{variations-IR, 173 | author = {Neil Ghani and 174 | Conor McBride and 175 | Fredrik Nordvall Forsberg and 176 | Stephan Spahn}, 177 | editor = {Kim G. Larsen and 178 | Hans L. Bodlaender and 179 | Jean{-}Fran{\c{c}}ois Raskin}, 180 | title = {Variations on Inductive-Recursive Definitions}, 181 | booktitle = {42nd International Symposium on Mathematical Foundations of Computer 182 | Science, {MFCS} 2017, August 21-25, 2017 - Aalborg, Denmark}, 183 | series = {LIPIcs}, 184 | volume = {83}, 185 | pages = {63:1--63:13}, 186 | publisher = {Schloss Dagstuhl - Leibniz-Zentrum f{\"{u}}r Informatik}, 187 | year = {2017}, 188 | url = {https://doi.org/10.4230/LIPIcs.MFCS.2017.63}, 189 | doi = {10.4230/LIPIcs.MFCS.2017.63}, 190 | timestamp = {Tue, 11 Feb 2020 15:52:14 +0100}, 191 | biburl = {https://dblp.org/rec/conf/mfcs/GhaniMFS17.bib}, 192 | bibsource = {dblp computer science bibliography, https://dblp.org}, 193 | annotation = "Explores non-closure under composition of Dybjer-Setzer IR codes, and Irish codes which are closed under composition." 194 | } 195 | 196 | @inproceedings{observational-type-theory, 197 | author = {Thorsten Altenkirch and 198 | Conor McBride and 199 | Wouter Swierstra}, 200 | editor = {Aaron Stump and 201 | Hongwei Xi}, 202 | title = {Observational equality, now!}, 203 | booktitle = {Proceedings of the {ACM} Workshop Programming Languages meets Program 204 | Verification, {PLPV} 2007, Freiburg, Germany, October 5, 2007}, 205 | pages = {57--68}, 206 | publisher = {{ACM}}, 207 | year = {2007}, 208 | url = {https://doi.org/10.1145/1292597.1292608}, 209 | doi = {10.1145/1292597.1292608}, 210 | timestamp = {Sun, 02 Jun 2019 21:19:59 +0200}, 211 | biburl = {https://dblp.org/rec/conf/plpv/AltenkirchMS07.bib}, 212 | bibsource = {dblp computer science bibliography, https://dblp.org}, 213 | annotation = "Presents Observational Type Theory.", 214 | } 215 | 216 | @book{hott-book, 217 | author = {The Univalent Foundations Program}, 218 | title = {Homotopy Type Theory: Univalent Foundations of Mathematics}, 219 | publisher = {Institute for Advanced Study}, 220 | year = {2013}, 221 | url = {https://homotopytypetheory.org/book/}, 222 | timestamp = {Fri, 06 Dec 2019 15:25:34 +0100}, 223 | biburl = {https://dblp.org/rec/books/daglib/0046165.bib}, 224 | bibsource = {dblp computer science bibliography, https://dblp.org}, 225 | annotation = "The HoTT book.", 226 | } 227 | 228 | @book{girard-proofs-and-types, 229 | author = "Jean-Yves Girard", 230 | title = "Proofs and Types", 231 | publisher = "Cambridge University Press", 232 | year = {1990}, 233 | url = {http://www.paultaylor.eu/stable/prot.pdf}, 234 | note = "Translated and with appendices by Paul Taylor and Yves Lafont.", 235 | annotation = "Talks about commutative conversions.", 236 | } 237 | 238 | @software{coq-8.12, 239 | author = {The Coq Development Team}, 240 | title = {The {C}oq Proof Assistant, version 8.12.0}, 241 | month = jul, 242 | year = 2020, 243 | publisher = {Zenodo}, 244 | version = {8.12.0}, 245 | doi = {10.5281/zenodo.4021912}, 246 | url = {https://doi.org/10.5281/zenodo.4021912}, 247 | } 248 | -------------------------------------------------------------------------------- /paper/whynotw.tex: -------------------------------------------------------------------------------- 1 | 2 | \documentclass[a4paper,UKenglish,cleveref,nameinlink,autoref,thm-restate]{lipics-v2019} 3 | %This is a template for producing LIPIcs articles. 4 | %See lipics-manual.pdf for further information. 5 | %for A4 paper format use option "a4paper", for US-letter use option "letterpaper" 6 | %for british hyphenation rules use option "UKenglish", for american hyphenation rules use option "USenglish" 7 | %for section-numbered lemmas etc., use "numberwithinsect" 8 | %for enabling cleveref support, use "cleveref" 9 | %for enabling autoref support, use "autoref" 10 | %for anonymousing the authors (e.g. for double-blind review), add "anonymous" 11 | %for enabling thm-restate support, use "thm-restate" 12 | 13 | \bibliographystyle{plainurl}% the mandatory bibstyle 14 | 15 | \title{Why not W?} 16 | 17 | \author{Jasper Hugunin}{Carnegie Mellon University, Pittsburgh PA, USA}{jasper@hugunin.net}{https://orcid.org/0000-0002-1133-5354}{} 18 | 19 | \authorrunning{J. Hugunin} 20 | 21 | \Copyright{Jasper Hugunin} 22 | 23 | \ccsdesc[500]{Theory of computation~Type theory} 24 | 25 | \keywords{dependent types, intensional type theory, inductive types, W types} 26 | 27 | \relatedversion{} %optional, e.g. full version hosted on arXiv, HAL, or other respository/website 28 | %\relatedversion{A full version of the paper is available at \url{...}.} 29 | 30 | \supplement{Tag \texttt{v0.1} of \url{https://github.com/jashug/WhyNotW}} 31 | 32 | \acknowledgements{I want to thank Jon Sterling and the anonymous reviewers for their helpful feedback.} 33 | 34 | %\nolinenumbers %uncomment to disable line numbering 35 | 36 | %\hideLIPIcs %uncomment to remove references to LIPIcs series (logo, DOI, ...), e.g. when preparing a pre-final version to be uploaded to arXiv or another public repository 37 | 38 | %Editor-only macros:: begin (do not touch as author)%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% 39 | \EventEditors{John Q. Open and Joan R. Access} 40 | \EventNoEds{2} 41 | \EventLongTitle{42nd Conference on Very Important Topics (CVIT 2016)} 42 | \EventShortTitle{CVIT 2016} 43 | \EventAcronym{CVIT} 44 | \EventYear{2016} 45 | \EventDate{December 24--27, 2016} 46 | \EventLocation{Little Whinging, United Kingdom} 47 | \EventLogo{} 48 | \SeriesVolume{42} 49 | \ArticleNo{23} 50 | %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% 51 | 52 | %\usepackage{float} 53 | %\floatstyle{boxed} 54 | %\restylefloat{figure} 55 | 56 | \newcommand{\zero}{\mathtt{0}} 57 | \newcommand{\one}{\mathtt{1}} 58 | \newcommand{\bool}{\mathtt{2}} 59 | \newcommand{\true}{\mathtt{tt}} 60 | \newcommand{\false}{\mathtt{ff}} 61 | \newcommand{\codeO}{\mathtt{\hat{O}}} 62 | \newcommand{\codeS}{\mathtt{\hat{S}}} 63 | \newcommand{\U}{\mathtt{U}} 64 | 65 | \DeclareMathOperator{\supop}{\mathtt{sup}} 66 | \renewcommand{\sup}[2]{\supop {#1}\:\!{#2}} 67 | 68 | \newcommand{\N}{\hyperref[define-N]{\mathbb{N}}} 69 | \newcommand{\preN}{\hyperref[define-preN]{\tilde{\mathbb{N}}}} 70 | 71 | \DeclareMathOperator{\case}{\mathtt{case}} 72 | \newcommand{\caset}[2]{\case {#1}\;\mathtt{of}\;\{{#2}\}} 73 | 74 | %\newcommand{\casepath}[4]{\case {#1};\mathtt{inl}\;\Id{\mathunderscore}{#2}\;\mathtt{return}\;{#3}\;\mathtt{of}\;\{\refl\mapsto{#4}\}} 75 | 76 | \DeclareMathOperator{\W}{\mathtt{W}} 77 | 78 | \DeclareMathOperator{\Idop}{\mathtt{Id}} 79 | \newcommand{\Id}[2]{\Idop {#1}\;{#2}} 80 | \newcommand{\IdA}[3]{\Idop_{#1}{#2}\;{#3}} 81 | \newcommand{\refl}{\mathtt{refl}} 82 | 83 | \DeclareMathOperator{\inl}{\mathtt{inl}} 84 | \DeclareMathOperator{\inr}{\mathtt{inr}} 85 | 86 | \DeclareMathOperator{\fst}{\mathtt{fst}} 87 | \DeclareMathOperator{\snd}{\mathtt{snd}} 88 | 89 | \DeclareMathOperator{\canonical}{\hyperref[define-canonical]{\mathtt{canonical}}} 90 | 91 | \newcommand{\ISO}{\hyperref[define-ISO]{\mathrm{ISO}}} 92 | \newcommand{\ISS}{\hyperref[define-ISS]{\mathrm{ISS}}} 93 | \newcommand{\IS}{\mathrm{IS}} 94 | 95 | \newcommand{\zerO}{\hyperref[define-O]{\mathrm{O}}} 96 | \DeclareMathOperator{\Succ}{\hyperref[define-S]{S}} 97 | \DeclareMathOperator{\recN}{\hyperref[define-recN]{\mathrm{rec\mathbb{N}}}} 98 | 99 | \DeclareMathOperator{\nonind}{nonind} 100 | \DeclareMathOperator{\ind}{ind} 101 | \DeclareMathOperator{\nil}{nil} 102 | \DeclareMathOperator{\choice}{choice} 103 | \newcommand{\Code}{\mathrm{Code}} 104 | 105 | \newcommand{\Ix}{\mathrm{Ix}} 106 | 107 | \DeclareMathOperator{\El}{El} 108 | \DeclareMathOperator{\preEl}{\tilde{El}} 109 | \DeclareMathOperator{\intro}{intro} 110 | \DeclareMathOperator{\rec}{rec} 111 | 112 | \DeclareMathOperator{\All}{All} 113 | 114 | \newcommand{\splitequiv}{\hyperref[split-equiv]{r}} 115 | 116 | \newcommand{\mathquote}[1]{\text{``${#1}$''}} 117 | 118 | \DeclareMathOperator{\funext}{\mathtt{funext}} 119 | 120 | \newcommand{\funF}{\hyperref[define-F]{F}} 121 | \newcommand{\funG}{\hyperref[define-G]{G}} 122 | 123 | \begin{document} 124 | 125 | \maketitle 126 | 127 | \begin{abstract} 128 | In an extensional setting, $\W$ types are sufficient to construct a broad class of inductive types, but in intensional type theory the standard construction of even the natural numbers does not satisfy the required induction principle. In this paper, we show how to refine the standard construction of inductive types such that the induction principle is provable and computes as expected in intensional type theory without using function extensionality. We extend this by constructing from $\W$ an internal universe of codes for inductive types, such that this universe is itself an inductive type described by a code in the next larger universe. We use this universe to mechanize and internalize our refined construction. 129 | \end{abstract} 130 | 131 | \section{Introduction} 132 | 133 | In intensional type theory with only type formers $\zero$, $\one$, $\bool$, $\Sigma$, $\Pi$, $\W$, $\Idop$ and $\U$, can the natural numbers be constructed? 134 | 135 | The $\W$ type \cite{mltt} captures the essence of induction (that we have a collection of possible cases, and for each case there is a collection of sub-cases to be handled inductively), and in extensional type theory it is straightforward to construct familiar inductive types out of it, including the natural numbers \cite{dybjer-W-encodes-inductive}. 136 | Taking the elements of the two-element type $\bool$ to be $\codeO$ and $\codeS$, we define \begin{equation}\label{define-preN}\preN = \W_{b : \bool} (\caset{b}{\codeO \mapsto \zero, \codeS \mapsto \one}).\end{equation} 137 | (the tilde distinguishes the standard construction from our refined construction of the natural numbers in \cref{construct-N}) 138 | 139 | However, as is well known \cite{dybjer-W-encodes-inductive,luo-wellordering,mcbride_wtypes,programming-in-MLTT}, in intensional type theory we cannot prove the induction principle for $\preN$ without some form of function extensionality. 140 | The obstacle is in the $\codeO$ case, where we end up needing to prove $P\;f$ for an arbitrary $f : \zero \to \preN$, when we only know $P\;(x\mapsto\caset{x}{})$. 141 | 142 | Can this obstacle be avoided? 143 | The answer turns out to be yes; in this paper, we show that refining the standard construction allows the natural numbers and many other inductive types to be constructed from $\W$ in intensional type theory. 144 | \footnote{These results have been formalized in Coq 8.12 \cite{coq-8.12}: see the link to supplementary material in the top matter of this article.} 145 | 146 | \paragraph*{Type-theoretic notations and assumptions} 147 | 148 | We work in a standard intensional type theory with 149 | dependent function types $\Pi_{a : A}B[a]$ (also written $\forall_{a : A}B[a]$, $(a : A) \to B[a]$, non-dependent version $A \to B$, constructed as $(x \mapsto y[x])$ or $(\lambda x.\;y[x])$), 150 | dependent pair types $\Sigma_{a : A}B[a]$ (also written $(a : A) \times B[a]$, non-dependent version $A \times B$, constructed as $(x, y)$, destructed as $\fst p$, $\snd p$), 151 | finite types $\zero$, $\one$ (with inhabitant $\star$), $\bool$ (with inhabitants $\false$ and $\true$, aliased to $\codeO$ and $\codeS$ when we are talking about constructing the natural numbers), 152 | $\W$ types $\W_{a : A}B[a]$ (constructor $\sup{a}{f}$ for $a : A$ and $f : B[a] \to \W_{a}B[a]$), 153 | identity types $\IdA{A}{x}{y}$ (constructor $\refl$, destruction of $e : \Id{x}{y}$ keeps $x$ fixed and generalizes over $y$ and $e$), 154 | and a universe $\U$. 155 | We define the coproduct $A + B$ as $\sum_{b : \bool}\caset{b}{\false\mapsto A, \true\mapsto B}$, and notate the injections as $\inl$ and $\inr$. 156 | 157 | Function extensionality is the principle that $\forall_x\Id{(f\;x)}{(g\;x)}$ implies $\Id{f}{g}$, and uniqueness of identity proofs is the principle that $\IdA{\Id{x}{y}}{p}{q}$ is always inhabited. We do \emph{not} assume either of these principles. 158 | 159 | We require strict $\beta$-rules for all type formers, and strict $\eta$ for $\Sigma$ (that $p = (\fst p, \snd p)$) and $\Pi$ (that $f = (x \mapsto f x)$). For convenience we will also assume strict $\eta$ for $\one$ (that $u = \star$). 160 | 161 | \section{Constructing \texorpdfstring{$\N$}{ℕ} (for real this time)}\label{construct-N} 162 | 163 | We run into problems in the $\codeO$ case because we don't know that $f = (x\mapsto\caset{x}{})$ for an arbitrary $f : \zero \to \preN$. 164 | To solve those problems, we will assume them away. 165 | To construct $\N$, we will first define a predicate $\canonical : \preN \to U$ such that $\canonical(\sup{\codeO}{f})$ implies $\Id{(x\mapsto\caset{x}{})}{f}$. 166 | We then let $\N = \Sigma_{x : \preN}\canonical x$ be the canonical elements of $\preN$ (with $\preN$ defined by \cref{define-preN}). 167 | This predicate will be defined by induction on $\W$, so we can start out with \[\canonical (\sup{x}{f}) =\; ? : \U\qquad\text{($x : \bool$, $f : \dots \to \preN$, may use $\canonical (f\; i) : \U$)}.\] 168 | The obvious next thing to do is to split by cases on $x : \bool$: 169 | \begin{gather*} 170 | \canonical (\sup{\codeO}{f}) =\; ? : \U\qquad\text{($f : \zero \to \preN$, may use $\canonical(f\;i)$)},\\ 171 | \canonical (\sup{\codeS}{f}) =\; ? : \U\qquad\text{($f : \one \to \preN$, may use $\canonical(f\;i)$)}. 172 | \end{gather*} 173 | 174 | We need canonical terms to be \emph{hereditarily} canonical, that is, we want to include the condition that all sub-terms are canonical. 175 | For the $\codeS$ case, thanks to the strict $\eta$ rules for $\one$ and $\Pi$, the types $\canonical(f\;\star)$ and $(i : \one) \to \canonical(f\;i)$ are equivalent; we can use either one. 176 | This will be the only condition we need for the $\codeS$ case, so we can complete this part of the definition: 177 | \[\canonical (\sup{\codeS}{f}) =\canonical(f\;\star).\] 178 | 179 | The $\codeO$ case is the interesting one. 180 | The blind translation of ``every sub-term is canonical'' is $(i : \zero) \to \canonical(f\;i)$, 181 | but this leads to the same problem as before: without function extensionality we can't work with functions out of $\zero$. 182 | Luckily, we have escaped the rigid constraints of the $\W$ type former, and have the freedom to translate the recursive condition as simply $\one$. 183 | No sub-terms of zero, no conditions necessary! 184 | \[\canonical (\sup{\codeO}{f}) =\; ? : \U\qquad\text{($f : \zero \to \preN$)}\] 185 | 186 | That is all well and good, but we can't forget why we are here in the first place: we need $\Id{(x\mapsto\caset{x}{})}{f}$. 187 | Luckily, there is a hole just waiting to be filled: 188 | \[\canonical (\sup{\codeO}{f}) = \Id{(x\mapsto\caset{x}{})}{f}.\] 189 | 190 | \begin{figure}[t] 191 | \fbox{\begin{minipage}{\textwidth}\setlength{\abovedisplayskip}{0pt} % Remove excess vertical space 192 | \begin{gather} 193 | \preN = \W_{b : \bool} (\caset{b}{\codeO \mapsto \zero, \codeS \mapsto \one}) : \U,\nonumber\\ 194 | \begin{aligned} 195 | \canonical& : \preN \to \U,\\ 196 | \canonical& (\sup{\codeO}{f}) = \Id{(x\mapsto\caset{x}{})}{f},\\ 197 | \canonical& (\sup{\codeS}{f}) = \canonical(f\;\star), 198 | \end{aligned}\label{define-canonical}\\ 199 | \N = \Sigma_{x : \preN}\canonical x : \U,\label{define-N}\\ 200 | \zerO = (\sup{\codeO}{(x\mapsto\caset{x}{})},\refl) : \N,\label{define-O}\\ 201 | \Succ = n \mapsto (\sup{\codeS}{(\star \mapsto \fst n)},\snd n) : \N \to \N.\label{define-S} 202 | \end{gather} 203 | \end{minipage}} 204 | \caption{The complete definition of $\N$.} 205 | \end{figure} 206 | 207 | \paragraph*{Induction} 208 | 209 | Now we are ready for the finale: induction for $\N$ with the right computational behavior. 210 | 211 | Assume we are given a type $P[n]$ which depends on $n : \N$, along with terms $\ISO : P[\zerO]$ and $\ISS : \forall_{n : \N}P[n] \to P[\Succ\;n]$. Our mission is to define a term $\recN : \forall_{n : \N}P[n]$. Happily, the proof goes through if we simply follow our nose. 212 | 213 | We begin by performing induction on $\fst n : \preN$, and then case on $\codeO$ vs $\codeS$, just like the definition of $\canonical$. 214 | \begin{gather*} 215 | \recN (\sup{\codeO}{f},y) =\; ? : P[(\sup{\codeO}{f},y)]\qquad\text{($f : \zero \to \preN$, $y : \Id{(x\mapsto\caset{x}{})}{f}$)},\\ 216 | \recN (\sup{\codeS}{f},y) =\; ? : P[(\sup{\codeS}{f},y)]\qquad\text{($f : \one \to \preN$, $y : \canonical(f\;\star)$)}.\\ 217 | \text{(where we may make recursive calls $\recN(f\; i, y')$ for any $i$ and $y'$)} 218 | \end{gather*} 219 | 220 | In the $\codeS$ case, $f = (\star \mapsto f\; \star)$ by the $\eta$ rules for $\one$ and $\Pi$, and thus $(\sup{\codeS}{f}, y) = \Succ\;(f\; \star, y)$. We can thus define 221 | \begin{equation*}\recN (\sup{\codeS}{f},y) = \ISS\;(f\;\star, y)\;(\recN(f\;\star, y)).\end{equation*} 222 | 223 | The $\codeO$ case is again the interesting one, but it is only a little tricky. We know $\ISO : P[ (\sup{\codeO}{(x\mapsto\caset{x}{})},\refl)]$, and we want $P[(\sup{\codeO}{f},y)]$. But since we have $y : \Id{(x\mapsto\caset{x}{})}{f}$, this is a direct application of the eliminator for $\Idop$. 224 | We thus complete the definition of $\recN$ with 225 | \begin{equation*}\recN (\sup{\codeO}{f},y) = \caset{y}{\refl \mapsto \ISO}.\end{equation*} 226 | 227 | Examining the definitions, we can see that as long as we have strict $\eta$ for $\Sigma$ and strict $\beta$ for $\Idop$, $\recN \zerO = \ISO$ and $\recN(\Succ n) = \ISS\;n\;(\recN\;n)$. Thus we have indeed defined the natural numbers with the expected induction principle and computational behavior in terms of the $\W$ type. 228 | 229 | \begin{figure}[t] 230 | \fbox{\begin{minipage}{\textwidth} 231 | Given 232 | \begin{gather} 233 | \text{a type $P[n]$ depending on $n : \N$},\\ 234 | \ISO : P[\zerO],\label{define-ISO}\\ 235 | \ISS : \forall_{n : \N}P[n] \to P[\Succ\;n],\label{define-ISS} 236 | \end{gather} 237 | we have 238 | \begin{gather} 239 | \begin{aligned} 240 | \recN& : \forall_{n : \N}P[n],\\ 241 | \recN& (\sup{\codeO}{f},y) = \caset{y}{\refl \mapsto \ISO},\\ 242 | \recN& (\sup{\codeS}{f},y) = \ISS\;(f\;\star, y)\;(\recN(f\;\star, y)), 243 | \end{aligned}\label{define-recN}\\ 244 | \recN \zerO = \ISO\label{recN-eqO},\\ 245 | \recN(\Succ n) = \ISS\;n\;(\recN\;n)\label{recN-eqS}. 246 | \end{gather} 247 | \end{minipage}} 248 | \caption{Induction for $\N$.} 249 | \end{figure} 250 | 251 | \begin{theorem} 252 | The natural numbers can be constructed in intensional type theory with only type formers $\zero$, $\one$, $\bool$, $\Sigma$, $\Pi$, $\W$, $\Idop$ and $\U$, such that the induction principle has the expected computational behavior. 253 | \end{theorem} 254 | 255 | \section{The General Case}\label{general-case} 256 | Above, we have refuted a widely held intuition about the expressiveness of intensional type theory with $\W$ as the only primitive inductive type. Once we know we can construct the natural numbers, that we can construct lots of other inductive types is much less surprising. 257 | 258 | Nevertheless, for completeness we define below an internal type of codes for inductive types along with the construction from $\W$ types of the interpretation of those codes. For convenience, in this section we assume that we have not just one universe $\U$ but an infinite cumulative tower of universes $\U_0 : \U_1 : \dots : \U_i : \U_{i+1} : \dots$ all closed under $\zero$, $\one$, $\bool$, $\Sigma$, $\Pi$, $\W$, and $\Idop$ such that $A : \U_i$ implies $A : \U_{i+1}$. 259 | 260 | The end result is a universe of inductive types which is self-describing, or ``levitating'' in the sense of \cite{levitation}. 261 | 262 | \subsection{Inductive Codes}\label{inductive-codes} 263 | We will let $\Code_i : \U_{i+1}$ be the type of codes for inductive types in $\U_i$, and implement it for now as a primitive inductive type. In \cref{bootstrap} we will show how to construct $\Code$ itself from $\W$. 264 | 265 | To define $\Code$, we adapt the axiomatization of induction-recursion from \cite{finite-axiom-IR}. Thus $\Code_i$ is generated by the constructors 266 | \[\nil : \Code_i,\quad \nonind : (A : \U_i) \to (A \to \Code_i) \to \Code_i, \quad \ind : \U_i \to \Code_i \to \Code_i.\] 267 | Looking at $\U_i$ as the usual category of types and functions, a code $A : \Code_i$ defines an endofunctor $\funF_A : \U_i \to \U_i$ defined by recursion on $A$ by 268 | \begin{gather} 269 | F_{\nil}\;X = \one,\label{define-F}\\ 270 | F_{\nonind(A,B)}\;X = \Sigma_{a : A} F_{(B\;a)}\; X,\\ 271 | F_{\ind(\Ix, B)}\;X = (\Ix \to X) \times F_{B}\; X. 272 | \end{gather} 273 | 274 | \begin{example}We can define a code for the natural numbers as \[\mathquote{\N} = \nonind(2, b \mapsto \caset{b}{\codeO \mapsto \nil, \codeS\mapsto \ind(1, \nil)}) : \Code_0.\] 275 | \end{example} 276 | 277 | Each code also defines a polynomial functor $G_A\;X = \Sigma_{s : S_A} (P_A\;s \to X)$, which is what is used in the standard construction: 278 | \begin{align} 279 | &S_{\nil} = 1 & &P_{\nil}\;\star = 0\label{define-G}\\ 280 | &S_{\nonind(A,B)} = \Sigma_{a : A} S_{(B\;a)} & &P_{\nonind(A, B)}\;(a,b) = P_{(B\;a)}\;b \\ 281 | &S_{\ind(\Ix, B)} = S_B & &P_{\ind(\Ix, B)}\;b = \Ix + P_B\;b.\\ 282 | &G_A\;X = \Sigma_{s : S_A} (P_A\;s \to X) & &\preEl A = \W_{s : S_A}P_A. 283 | \end{align} 284 | The idea here is that $S_A$ collects up all the non-inductive data, and then $P_A$ counts the number of inductive sub-cases. 285 | 286 | There is an easy-to-define natural transformation $\epsilon : \funF \Rightarrow \funG$, and it even has a left inverse on objects, but without function extensionality $\epsilon$ does not have a right inverse (roughly speaking, $\epsilon$ is not surjective); there are usually terms $g : \funG\;X$ not in the image of $\epsilon$. This is exactly the problem we ran into in the case of the natural numbers: the map $(\star \mapsto (x \mapsto \caset{x}{})) : 1 \to (0 \to X)$ is not surjective. 287 | (The above $S$, $P$, and $\epsilon$ roughly correspond to Lemma 3 in \cite{dybjer-W-encodes-inductive}) 288 | 289 | The last component we need is $\All_A\;s : (Q : P_A\;s \to \U_j) \to \U_j$ (for universe level $j\geq i$), the quantifier ``holds at every position'' (a refinement of $\forall_p, Q\;p$): 290 | \begin{gather} 291 | \All_{\nil}\star\;Q = \one,\\ 292 | \All_{\nonind(A,B)}(a,b)\;Q = \All_{(B\;a)}b\;Q,\\ 293 | \All_{\ind(\Ix, B)}b\;Q = (\forall_i,Q\;(\inl i)) \times \All_{B}b\;(Q \circ \inr). 294 | \end{gather} 295 | 296 | Noting that $\snd (\epsilon\; t) : P\;(\fst (\epsilon\;t)) \to X$ enumerates the sub-terms of $t : \funF\;X$, we find that $\All (Q \circ \snd (\epsilon\;t))$ lifts a predicate $Q : X \to \U_j$ to a predicate over $t : \funF\;X$. 297 | 298 | \begin{lemma}\label{split-equiv}There is an equivalence $\splitequiv$ (\`a la Voevodsky, a function with contractible fibers) \begin{equation} 299 | \splitequiv : \funF\;(\Sigma_{x : X}C\;x) \simeq \Sigma_{(t : \funF\;X)}\All (C\circ \snd (\epsilon\;t)). 300 | \end{equation} 301 | \end{lemma} 302 | 303 | \begin{proof} 304 | Follows easily by induction on the code $A$. We use equivalences \`a la Voevodsky as a concrete definition of coherent equivalences, which are the ``right'' way to define type equivalence in the absence of UIP. 305 | \end{proof} 306 | 307 | \subsection{The General Construction} 308 | We are finally ready to define the true construction of inductive types $\El : \Code \to \U_i$. As with natural numbers, we define a ``canonicity'' predicate on $\preEl A$, which says that ``all subterms are canonical, and this node is in the image of $\epsilon$''. 309 | This translates as: 310 | \begin{equation} 311 | \canonical (\sup{s}{f}) = \All (\canonical \circ f) \times (t : \funF\;(\preEl A)) \times \IdA{\funG\;(\preEl\;A)}{(\epsilon\; t)}{(s, f)} : \U_i,\end{equation} 312 | and thus we finally have \begin{equation}\El\;A = \Sigma_{x : \preEl A} \canonical x.\end{equation} 313 | For the constructors, we expect to have $\intro : \funF\;(\El A) \to \El A$, which we define by 314 | \begin{equation} 315 | \intro x = (\supop\;(\epsilon\;(\fst\;(\splitequiv\;x))),(\snd(\splitequiv\;x),\fst\;(\splitequiv\;x),\refl)). 316 | \end{equation} 317 | using the equivalence $\splitequiv$ from \cref{split-equiv} to split $x : \funF\;(\El A)$ into $\fst(\splitequiv\;x) : \funF\;(\preEl A)$ and $\snd(\splitequiv\;x) : \All(\canonical \circ \snd(\epsilon\;\fst\;(\splitequiv\;x)))$. 318 | 319 | \subsection{General Induction} 320 | When we go to define the induction principle for $\El A$, we are given $P : \El A \to \U_j$ for some $j \geq i$ and the induction step $\IS : \forall_{(x : F\;(\El A))}\All(P \circ \snd(\epsilon\;x)) \to P\;(\intro x)$, and want to define $\rec : \forall_{(x : \El A)}P\;x$. The definition proceeds by induction on $\fst x$: 321 | \[\rec(\sup{s}{f},(h,t,e)) ={} ? : P\;(\sup{s}{f},(h,t,e))\qquad h : \All(\canonical \circ f) \quad e : \Id{(\epsilon \;t)}{(s,f)},\] and we have induction hypothesis $H = p \mapsto c \mapsto \rec(f\;p,c) : \Pi_p\Pi_c P\;(f\;p,c)$. Next, we destruct the identity proof $e$, generalizing over both $h$ and $H$, leaving us with 322 | \[\rec(\supop(\epsilon\;t),(h,t,\refl)) ={} ? : P\;(\supop(\epsilon\;t),(h,t,\refl)),\] for $t : \funF\;(\preEl A)$, $h : \All(\canonical\circ\snd(\epsilon\;t))$, and $H : \Pi_p\Pi_c P\;(\snd(\epsilon\;t)\;p,c)$. The last step to bring us in line with the definition of $\intro$ is to use the equivalence from \cref{split-equiv} to replace $(t, h)$ with $\splitequiv\;x$ for some $x : \funF\;(\El A)$, leaving us with 323 | \[\rec(\supop(\epsilon\;(\fst(\splitequiv\;x))),(\snd(\splitequiv\;x),\fst(\splitequiv\;x),\refl)) ={} ? : P\;(\intro x)\] 324 | and induction hypothesis $H : \Pi_p\Pi_c P\;(\snd(\epsilon\;(\fst(\splitequiv\;x)))\;p,c)$. We can then apply $\IS$, but that leaves us with an obligation to prove $\All (P\circ\snd(\epsilon\;x))$. Fortunately, it is easy to show by induction on the code $A$ that our hypothesis $H$ is sufficient to dispatch this obligation. 325 | 326 | This completes the definition of the induction principle, and it can be observed on concrete examples like the natural numbers to have the expected computational behavior. We can also prove a propositional equality $\Id{(\rec(\intro x))}{(\IS\;x\;(\rec\circ \snd(\epsilon\;x)))}$ witnessing the expected computation rule, and observe on concrete examples that this witness computes to reflexivity. The details of this construction have all been formalized in Coq. 327 | 328 | \subsection{Bootstrapping}\label{bootstrap} 329 | In \cref{inductive-codes} we postulated the type $\Code_i$ to be a primitive inductive type, which leads to the question of whether the general construction we have proposed is \emph{really} constructing inductive types out of $\W$ or whether it is making sneaky use of the inductive structure of $\Code_i$ to perform the construction. 330 | 331 | As a first observation, $\Code_i : \U_{i+1}$ while $\El : \Code_i \to \U_i$, thus $\Code_i$ can't appear as data in $\El A$: it is too big! However, this argument doesn't show that we can completely eliminate $\Code_i$ from the construction. 332 | 333 | Next, we observe that the inductive type $\Code_i$ itself has a code $\mathquote{\Code_i} : \Code_{i+1}$: 334 | \begin{align*} 335 | \mathquote{\Code_i} ={}& \nonind((1+\U_i)+\U_i,t \mapsto \case t\;\mathtt{of}\;\{\\ 336 | &\quad \inl(\inl\star) \mapsto \nil,&&\text{(case $\nil$)} \notag\\ 337 | &\quad \inl(\inr A) \mapsto \ind(A,\nil),&&\text{(case $\nonind$)} \notag\\ 338 | &\quad \inr \Ix \mapsto \ind(1,\nil),&&\text{(case $\ind$)} \notag\\ 339 | &\notag\}). 340 | \end{align*} 341 | 342 | Then we can propose to define $\Code_i = \El \mathquote{\Code_i}$, but this is a circular definition: we define $\Code_i$ by using recursion on $\Code_{i+1}$. What we really want, and in some ways should be able to expect, is that $\El\mathquote{\Code_i}$ \emph{computes} to a normal form which no longer mentions $\Code$ but is expressed purely in terms of $\W$. We could then tie the knot by defining $\Code_i$ to be what $\El \mathquote{\Code_i}$ \emph{will compute to}, once we have defined $\El$. 343 | 344 | There is just one minor, rather technical problem to resolve, which is that currently $\El$ (which is defined by recursion on codes) gets stuck on $\El (\caset{t}{\dots})$ which is used to branch on constructor tags; 345 | we are missing some sort of commuting conversion \cite[section 10]{girard-proofs-and-types}. 346 | Fortunately, this problem is easy to work around by reifying the operation of branching on constructor tags as part of $\Code$. We add another constructor 347 | \begin{equation} 348 | \choice : \Code_i \to \Code_i \to \Code_i,\qquad \funF_{\choice(A,B)}\;X = F_A\;X + F_B\;X 349 | \end{equation} 350 | which encodes the simple binary sum of functors, specializing the dependent sum of functors $\nonind(2,b\mapsto\caset{b}{\dots})$ (but with all proofs essentially the same). With this in hand, we can define 351 | \begin{align} 352 | \mathquote{\Code_i} ={}& \choice(\choice(\\ 353 | &\quad\nil,&\text{(case $\nil$)}\notag\\ 354 | &\quad\choice(\notag\\ 355 | &\qquad\nonind(\U_i,A\mapsto \ind(A,\nil)),&\text{(case $\nonind$)}\notag\\ 356 | &\qquad\ind(1,\ind(1,\nil)))),&\text{(case $\choice$)}\notag\\ 357 | &\quad\nonind(\U_i,\Ix \mapsto \ind(1,\nil))).&\text{(case $\ind$)}\notag 358 | \end{align} 359 | 360 | With this adjustment, the structure of the code is not hidden inside $\case$, and the computation of $\El\mathquote{\Code_i}$ proceeds to completion without becoming stuck, resulting in a term which does not mention $\Code$ at all. From there, we can define $\El$ such that $\El\mathquote{\Code_i} = \Code_i$, as in \cite{levitation} but with no invisible cables, just the $\W$ type. 361 | \begin{theorem} 362 | In intensional type theory with type formers $\zero$, $\one$, $\bool$, $\Sigma$, $\Pi$, $\W$, $\Idop$ and an infinite tower of universes $\U_i$, we can construct terms $\Code_i : \U_{i+1}$ and $\El : \Code_i \to \U_i$ such that $\El A$ is an inductive type, and we can also construct terms $\mathquote{\Code_i} : \Code_{i+1}$ such that $\El\mathquote{\Code_i} = \Code_i$. Furthermore, $\Code_i$ is not trivial: it contains codes for natural numbers, lists, binary trees, and many other inductive types, including inductive types such as $\W$ that have infinitary inductive arguments. 363 | \end{theorem} 364 | 365 | \section{Discussion} 366 | 367 | \subsection{Composition} 368 | Being codes for functors, one may ask if $\Code_i$ is closed under composition of functors? As with the codes for inductive-recursive types we have modified, without function extensionality we do not appear to have composition (for similar reasons as considered in \cite{variations-IR}). Indeed, experiments suggest that the general construction of a class of inductive types closed under composition of the underlying functors essentially requires function extensionality. Even worse, to get definitional computation rules for the resulting inductive types, all our attempts have required that transporting over $\funext(x\mapsto \refl)$ computes to the identity, a property which not even cubical type theory \cite{cchm-cubical} satisfies (it is satisfied, however, by observational type theory \cite{observational-type-theory}). Thus, we do not know how to combine a class of inductive types closed under composition constructed from the $\W$ type as we have in \cref{general-case} with the the principle of Univalence \cite{hott-book} while maintaining good computational behavior. 369 | 370 | We do however wish to emphasize that the construction in \cref{general-case} (which is not closed under composition) is completely compatible with Univalence, and could be implemented in cubical type theory as long as an identity type with strict $\beta$ rule is used. 371 | 372 | \subsection{Canonicity} 373 | Despite being constructed from $W$ types, our natural numbers enjoy the canonicity property (that for every closed term $n$ of type $\N$, either $n = \zerO$ or $n = \Succ m$ for some closed $m : \N$), at least as long as $\bool$ and $\Idop$ enjoy canonicity (closed $b : \bool$ implies $b = \codeO$ or $b = \codeS$, and closed $e : \Id{x}{y}$ implies $e = \refl$ and $x = y$). The trick is that when we have some representation of zero, it looks like $(\sup{\codeO}{f},e)$, where $e$ is a closed term of type $\Id{(x\mapsto\caset{x}{})}{f}$, and thus by canonicity for $\Idop$, this must be $(\sup{\codeO}{(x\mapsto\caset{x}{})},\refl) = \zerO$. 374 | 375 | However, in a situation like cubical type theory where function extensionality holds, $\Idop$ no longer enjoys canonicity, and neither does our construction of the natural numbers. 376 | 377 | \subsection{Problems} 378 | What are the problems with using this construction as the foundation for inductive types in a proof assistant? While we have shown bare possibility, this is not an obviously superior solution when compared to the inductive schemes present in proof assistants today. 379 | 380 | The construction is complex, which has the possibility of confusing unification and other elaboration algorithms. While the reduction behavior simulates the expected such, the reduction engine has to make many steps to simulate one step of a primitive inductive type, which can lead to a large slowdown. As an example, we observed the general construction slow down from seconds to check to half an hour when replacing primitive inductive types the bootstrapped definition of $\Code$. Understanding exactly why this slowdown happens and how to alleviate it is an important question to be answered before attempting to apply this construction in practice. 381 | 382 | There are also some (fairly esoteric) limitations to the expressivity of this construction. Nested inductive types such as $\mathtt{Inductive}\;\mathtt{tree} := \mathtt{node} : \mathtt{list}\;\mathtt{tree} \to \mathtt{tree}$ do not appear to be constructible, nor do mutual inductive types landing in a mixture of impredicative and predicative sorts at different levels, and nor do inductive-inductive types. 383 | 384 | \subsection{Setoids} 385 | In \cite{palmgren2019setoids}, Palmgren uses W types to construct a setoid model of extensional type theory in intensional type theory, including the natural numbers. In contrast, we have different goals (we are not concerned with extensional type theory), and our construction has different properties: we construct the natural numbers as a set not a setoid, with definitional computation rules and canonicity rather than working only up to an extensional setoid notion of equality. Other work on setoid models includes \cite{hoffman-thesis} and \cite{altenkirch-exteq}. 386 | 387 | \subsection{Conclusion} 388 | 389 | We have shown that intensional type theory with $\W$ and $\Idop$ types is more expressive than was previously believed. It supports not only the natural numbers, but a whole host of inductive types, generated by an internal type of codes, which is itself an inductive type coded for by itself (one universe level up). This brings possibilities for writing generic programs acting on inductive types internally (like in \cite{dybjer-generic}), and perhaps simplifies the general study of extensions of intensional type theory: once you know $\W$ works, you know lots of inductive types work. 390 | 391 | %Since intensional type theory forms the kernel of popular proof assistants such as Coq and Agda, understanding how to simplify intensional type theory has the potential to lead to simpler, more flexible, and/or easier to use proof assistants. 392 | 393 | Thus we return to the titular question: why not use $\W$ as the foundation of inductive types, for example in a proof assistant like Coq or Agda? Equipped with this result, one can no longer say that it is impossible. 394 | 395 | %% 396 | %% Bibliography 397 | %% 398 | 399 | %% Please use bibtex, 400 | 401 | \bibliography{whynotw} 402 | 403 | \appendix 404 | 405 | \end{document} 406 | --------------------------------------------------------------------------------