Skip to content

Commit 133d244

Browse files
authored
def: macros for creating copattern definitions (#376)
There are a couple of places in the 1Lab where we manually eta-expand out copattern definitions by hand to get better goal display. This mostly comes up when dealing with things like subcategories/forgetful functors; Agda will very happily unfold your nicely named category into `Restrict Blah`, which is not particularly helpful! Manually performing these copattern expansions is a bit of a pain, so this PR adds a small macro that performs this mechanical busywork for us. This removes the need for the `declare-concrete-category` macro in #375.
1 parent b423ee4 commit 133d244

File tree

16 files changed

+368
-119
lines changed

16 files changed

+368
-119
lines changed

src/1Lab/Reflection.lagda.md

Lines changed: 4 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -201,6 +201,10 @@ under-abs (lam v (abs nm _)) m = extend-context nm (arg (arginfo v (modality rel
201201
under-abs (pi a (abs nm _)) m = extend-context nm a m
202202
under-abs _ m = m
203203
204+
extend-context* : ∀ {a} {A : Type a} → Telescope → TC A → TC A
205+
extend-context* [] a = a
206+
extend-context* ((nm , tm) ∷ xs) a = extend-context nm tm (extend-context* xs a)
207+
204208
new-meta : Term → TC Term
205209
new-meta ty = do
206210
mv ← check-type unknown ty

src/1Lab/Reflection/Copattern.agda

Lines changed: 243 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,243 @@
1+
open import 1Lab.Reflection.Signature
2+
open import 1Lab.Reflection.Subst
3+
open import 1Lab.Reflection
4+
open import 1Lab.Path
5+
open import 1Lab.Type
6+
7+
module 1Lab.Reflection.Copattern where
8+
9+
--------------------------------------------------------------------------------
10+
-- Macros for manipulating copattern definitions.
11+
12+
-- Make a top-level copattern binding for an existing record.
13+
make-copattern : Bool Name Term Term TC ⊤
14+
make-copattern declare? def-name tm tp = do
15+
-- Ensure that codomain is a record type.
16+
let tele , cod-tp = pi-view tp
17+
def rec-name params pure cod-tp
18+
where _ typeError [ "make-copattern: codomain of " , termErr tp , " is not a record type." ]
19+
20+
let inst-tm = apply-tm* tm (tel→args 0 tele)
21+
22+
-- Construct copattern clauses for each field.
23+
ctor , fields get-record-type rec-name
24+
clauses
25+
in-context (reverse tele) $
26+
for fields λ (arg field-info field-name) do
27+
-- Infer the type of the field with all known arguments instantiated.
28+
field-tp infer-type (def field-name (argN inst-tm ∷ []))
29+
30+
-- Agda will insert implicits when defining copatterns even
31+
-- with 'withExpandLast true', so we need to do implicit instantiation
32+
-- by hand. There are also cases where it's better to fully
33+
-- eta-expand than not (e.g. the 'helper' we're expanding has a
34+
-- field defined by lazy matching, which does not reduce unless
35+
-- applied, and would cause duplication of the big input term). So
36+
-- we fully eta-expand clauses here.
37+
-- First, we strip off all leading quantifiers from the field
38+
-- type.
39+
let
40+
(field-tele , tp) = pi-view field-tp
41+
nargs = length field-tele
42+
clause-tele = tele ++ field-tele
43+
44+
-- Construct the pattern portion of the clause, making sure to
45+
-- bind all variables. Note that copattern projections are always
46+
-- visible.
47+
let
48+
pat = tel→pats nargs tele ++
49+
arg (set-visibility visible field-info) (proj field-name) ∷
50+
tel→pats 0 field-tele
51+
52+
-- Construct the body of the clause, making sure to apply all
53+
-- arguments bound outside the copattern match, and apply the
54+
-- eta-expanded arguments. We also need to apply all of the
55+
-- implicit arguments to 'tm'.
56+
body
57+
in-context (reverse clause-tele) $
58+
reduce (def field-name (raise nargs inst-tm v∷ tel→args 0 field-tele))
59+
60+
-- Construct the final clause.
61+
pure $ clause clause-tele pat body
62+
63+
-- Define a copattern binding, and predeclare its type if required.
64+
case declare? of λ where
65+
true declare (argN def-name) tp <|> pure tt
66+
false pure tt
67+
68+
-- Construct the final copattern.
69+
define-function def-name clauses
70+
71+
-- Repack a record type.
72+
-- Will also accept functions into record types like `A → Record`,
73+
-- and will perform a pointwise repacking.
74+
repack-record : Term Term TC Term
75+
repack-record tm tp = do
76+
-- Ensure that codomain is a record type.
77+
tele , cod-tp pi-view <$> reduce tp
78+
def rec-name params pure cod-tp
79+
where _ typeError [ "repack-record: codomain of " , termErr tp , " is not a record type." ]
80+
81+
-- Instantiate the term with all telescope arguments to saturate it.
82+
let inst-tm = apply-tm* tm (tel→args 0 tele)
83+
84+
-- Construct fields for each projection.
85+
ctor , fields get-record-type rec-name
86+
args
87+
in-context (reverse tele) $
88+
for fields λ (arg field-info field-name)
89+
argN <$> reduce (def field-name (argN inst-tm ∷ []))
90+
91+
-- Builld a pointwise repacking.
92+
pure (tel→lam tele (con ctor args))
93+
94+
-- Helper for the 'define' macros; Unifies the given goal with the type
95+
-- of the given function, if it has been defined. If the function has
96+
-- not been defined, and the first argument is 'false', then an error is
97+
-- raised.
98+
type-for : String Bool Name Term TC ⊤
99+
type-for tac decl? fun goal with decl?
100+
... | true = (unify goal =<< get-type fun) <|> pure tt
101+
... | false = (unify goal =<< get-type fun) <|> typeError
102+
[ "define-" , strErr tac , ": the function " , nameErr fun , " should already have been declared."
103+
]
104+
105+
--------------------------------------------------------------------------------
106+
-- Usage
107+
108+
{-
109+
Write the following in a module:
110+
> unquoteDecl copat = declare-copattern copat thing-to-be-expanded
111+
112+
If you wish to give the binding a type annotation, you can also use
113+
114+
> copat : Your-type
115+
> unquoteDecl copat = declare-copattern copat thing-to-be-expanded
116+
117+
Note that, in this case, the thing-to-be-expanded must have exactly the
118+
same type as the binding `copat`. All features of non-recursive records
119+
are supported, including instance fields and fields with implicit
120+
arguments.
121+
122+
These macros also allow you to lift functions 'A → some-record-type'
123+
into copattern definitions. Note that Agda will generate meta for
124+
implicits before performing quotation, so we need to explicitly
125+
bind all implicits using a lambda before quotation!
126+
-}
127+
128+
declare-copattern : {ℓ} {A : Type ℓ} Name A TC ⊤
129+
declare-copattern {A = A} nm x = do
130+
`x quoteTC x
131+
`A quoteTC A
132+
make-copattern true nm `x `A
133+
134+
define-copattern
135+
: {ℓ} (nm : Name)
136+
{@(tactic (type-for "copattern" true nm)) A : Type ℓ}
137+
A TC ⊤
138+
define-copattern nm {A = A} x = do
139+
`A quoteTC A
140+
`x define-abbrev nm "value" `A =<< quoteTC x
141+
make-copattern false nm `x `A
142+
143+
{-
144+
There is a slight caveat with level-polymorphic defintions, as
145+
they cannot be quoted into any `Type ℓ`. With this in mind,
146+
we also provide a pair of macros that work over `Typeω` instead.
147+
-}
148+
149+
declare-copatternω : {U : Typeω} Name U TC ⊤
150+
declare-copatternω nm A = do
151+
`A quoteωTC A
152+
-- Cannot quote things in type Typeω, but we can infer their type.
153+
`U infer-type `A
154+
make-copattern true nm `A `U
155+
156+
define-copatternω
157+
: (nm : Name) {@(tactic (type-for "copatternω" false nm)) U : Typeω}
158+
U TC ⊤
159+
define-copatternω nm A = do
160+
`U get-type nm
161+
`A define-abbrev nm "value" `U =<< quoteωTC A
162+
make-copattern false nm `A `U
163+
164+
{-
165+
Another common pattern is that two records `r : R p` and `s : R q` may contain
166+
the same data, but they are parameterized by different values.
167+
The `define-eta-expansion` macro will automatically construct a function
168+
`R p → R q` that eta-expands the first record out into a copattern definition.
169+
-}
170+
171+
define-eta-expansion : Name TC ⊤
172+
define-eta-expansion nm = do
173+
tp reduce =<< infer-type (def nm [])
174+
let tele , _ = pi-view tp
175+
let tm = tel→lam tele (var 0 [])
176+
make-copattern false nm tm tp
177+
178+
--------------------------------------------------------------------------------
179+
-- Tests
180+
181+
private module Test where
182+
-- Record type that uses all interesting features.
183+
record Record {ℓ} (A : Type ℓ) : Type ℓ where
184+
no-eta-equality
185+
constructor mk
186+
field
187+
⦃ c ⦄ : A
188+
{ f } : A A
189+
const : {x} f x ≡ c
190+
191+
-- Record created via a constructor.
192+
via-ctor : Record Nat
193+
via-ctor = mk ⦃ c = 0 ⦄ {f = λ _ 0} refl
194+
195+
-- Both macros work.
196+
unquoteDecl copat-decl-via-ctor = declare-copattern copat-decl-via-ctor via-ctor
197+
198+
copat-def-via-ctor : Record Nat
199+
unquoteDef copat-def-via-ctor = define-copattern copat-def-via-ctor via-ctor
200+
201+
-- Record created by a function.
202+
module _ (r : Record Nat) where
203+
open Record r
204+
via-function : Record Nat
205+
via-function .c = suc c
206+
via-function .f = suc ∘ f
207+
via-function .const = ap suc const
208+
209+
-- Also works when applied to the result of a function.
210+
unquoteDecl copat-decl-via-function = declare-copattern copat-decl-via-function (via-function via-ctor)
211+
212+
-- Test to see how we handle unused parameters.
213+
record Unused (n : Nat) : Type where
214+
field
215+
actual : Nat
216+
217+
zero-unused-param : Unused 0
218+
zero-unused-param = record { actual = 0 }
219+
220+
one-unused-param : {n} Unused n
221+
unquoteDef one-unused-param = declare-copattern one-unused-param zero-unused-param
222+
-- This is a type error:
223+
-- unquoteDef one-unused-param = define-copattern one-unused-param zero-unused-param
224+
-- because the 'define' macro propagates the type of the thing being
225+
-- defined inwards.
226+
227+
-- Functions into records that are universe polymorphic
228+
neat : {ℓ} {A : Type ℓ} A Record A
229+
neat a .Record.c = a
230+
neat a .Record.f _ = a
231+
neat a .Record.const = refl
232+
233+
-- Implicit insertion is correct for the define- macro, since the type
234+
-- of the function is given.
235+
cool : {ℓ} {A : Type ℓ} A Record A
236+
unquoteDef cool = define-copatternω cool neat
237+
238+
-- Eta-expanders
239+
expander : {m n : Nat} Unused m Unused n
240+
unquoteDef expander = define-eta-expansion expander
241+
242+
-- Raises a type error: the function should have a declaration.
243+
-- unquoteDecl uncool = define-copatternω uncool neat

src/1Lab/Reflection/Signature.agda

Lines changed: 12 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -211,6 +211,18 @@ helper-function def-nm suf ty cls = do
211211
_ define-function nm cls
212212
pure nm
213213

214+
-- Given a well-typed `val : ty`, return a definitionally-equal atomic
215+
-- term equal to `val`, potentially by lifting it into the signature.
216+
-- See 'helper-function' for the naming scheme.
217+
define-abbrev : Name String Term Term TC Term
218+
define-abbrev def-nm suf ty val with is-atomic-tree? val
219+
... | true = pure val
220+
... | false = do
221+
let (tel , _) = pi-impl-view ty
222+
nm helper-function def-nm suf ty
223+
[ clause tel (tel→pats 0 tel) (apply-tm* val (tel→args 0 tel)) ]
224+
pure (def₀ nm)
225+
214226
private
215227
make-args : Nat List (Arg Nat) List (Arg Term)
216228
make-args n xs = reverse $ map (λ (arg ai i) arg ai (var (n - i - 1) [])) xs

src/1Lab/Reflection/Subst.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -142,7 +142,7 @@ subst-tm ρ (agda-sort s) with s
142142
subst-tel ρ [] = []
143143
subst-tel ρ ((x , arg ai t) ∷ xs) = (x , arg ai (subst-tm ρ t)) ∷ subst-tel (liftS 1 ρ) xs
144144

145-
subst-clause σ (clause tel ps t) = clause (subst-tel σ tel) ps (subst-tm (wkS (length tel) σ) t)
145+
subst-clause σ (clause tel ps t) = clause (subst-tel σ tel) ps (subst-tm (liftS (length tel) σ) t)
146146
subst-clause σ (absurd-clause tel ps) = absurd-clause (subst-tel σ tel) ps
147147

148148
_<#>_ : Term Arg Term Term

src/Cat/Abelian/Images.lagda.md

Lines changed: 6 additions & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -53,8 +53,8 @@ images : ∀ {A B} (f : Hom A B) → Image C f
5353
images f = im where
5454
the-img : ↓Obj (const! (cut f)) Forget-full-subcat
5555
the-img .x = tt
56-
the-img .y .object = cut (Ker.kernel (Coker.coeq f))
57-
the-img .y .witness {c} = kernels-are-subobjects C ∅ _ (Ker.has-is-kernel _)
56+
the-img .y .fst = cut (Ker.kernel (Coker.coeq f))
57+
the-img .y .snd {c} = kernels-are-subobjects C ∅ _ (Ker.has-is-kernel _)
5858
```
5959
6060
Break $f$ down as an epi $p : A \epi \ker (\coker f)$ followed by a mono
@@ -128,7 +128,7 @@ a (unique) map $\coker (\ker f) \to X$ s.t. the triangle above commutes!
128128
where abstract
129129
path : other .map .map ∘ 0m ≡ other .map .map ∘ kernel f .Kernel.kernel
130130
path =
131-
other .y .witness _ _ $ sym $
131+
other .y .snd _ _ $ sym $
132132
pulll (other .map .commutes)
133133
·· Ker.equal f
134134
·· ∅.zero-∘r _
@@ -154,14 +154,14 @@ is the image of $f$.
154154
(coker-ker≃ker-coker f .is-invertible.invr))) refl
155155
∙ pullr (Coker.factors _) ∙ other .map .commutes)
156156
(sym (decompose f .snd ∙ assoc _ _ _))
157-
factor .sq = /-Hom-path $ sym $ other .y .witness _ _ $
157+
factor .sq = /-Hom-path $ sym $ other .y .snd _ _ $
158158
pulll (factor .β .commutes)
159159
∙ the-img .map .commutes
160160
·· sym (other .map .commutes)
161-
·· ap (other .y .object .map ∘_) (intror refl)
161+
·· ap (other .y .fst .map ∘_) (intror refl)
162162
163163
unique : ∀ x → factor ≡ x
164-
unique x = ↓Hom-path _ _ refl $ /-Hom-path $ other .y .witness _ _ $
164+
unique x = ↓Hom-path _ _ refl $ /-Hom-path $ other .y .snd _ _ $
165165
sym (x .β .commutes ∙ sym (factor .β .commutes))
166166
```
167167
</details>

src/Cat/Diagram/Equaliser/RegularMono.lagda.md

Lines changed: 7 additions & 7 deletions
Original file line numberDiff line numberDiff line change
@@ -43,7 +43,7 @@ is an [[equaliser]] of some pair of arrows $g, h : b \to c$.
4343
{c} : Ob
4444
arr₁ arr₂ : Hom b c
4545
has-is-eq : is-equaliser C arr₁ arr₂ f
46-
46+
4747
open is-equaliser has-is-eq public
4848
```
4949

@@ -53,7 +53,7 @@ are in fact monomorphisms:
5353
```agda
5454
is-regular-mono→is-mono : is-monic f
5555
is-regular-mono→is-mono = is-equaliser→is-monic _ has-is-eq
56-
56+
5757
open is-regular-mono using (is-regular-mono→is-mono) public
5858
```
5959

@@ -126,7 +126,7 @@ $\phi \circ i_2 = \rm{arr_2}$.
126126
```agda
127127
phi : Hom f⊔f.coapex reg.c
128128
phi = f⊔f.universal reg.equal
129-
129+
130130
open is-effective-mono
131131
mon : is-effective-mono C f
132132
mon .cokernel = f⊔f.coapex
@@ -201,20 +201,20 @@ the code below demonstrates.
201201
→ M-image C (is-regular-mono C , is-regular-mono→is-mono) f
202202
is-effective-mono→image {f = f} mon = im where
203203
module eff = is-effective-mono mon
204-
204+
205205
itself : ↓Obj _ _
206206
itself .x = tt
207-
itself .y = restrict (cut f) eff.is-effective-mono→is-regular-mono
207+
itself .y = cut f , eff.is-effective-mono→is-regular-mono
208208
itself .map = record { map = id ; commutes = idr _ }
209-
209+
210210
im : Initial _
211211
im .bot = itself
212212
im .has⊥ other = contr hom unique where
213213
hom : ↓Hom _ _ itself other
214214
hom .α = tt
215215
hom .β = other .map
216216
hom .sq = trivial!
217-
217+
218218
unique : ∀ x → hom ≡ x
219219
unique x = ↓Hom-path _ _ refl
220220
(ext (intror refl ∙ ap map (x .sq) ∙ elimr refl))

0 commit comments

Comments
 (0)