agda / agda

Agda is a dependently typed programming language / interactive theorem prover.

Home Page:https://wiki.portal.chalmers.se/agda/pmwiki.php

Geek Repo:Geek Repo

Github PK Tool:Github PK Tool

Module parameter `@++` annotations

lawcho opened this issue · comments

Whilst experimenting with #6385, I was surprised:

{-# OPTIONS --polarity #-}

open import Agda.Builtin.Unit
open import Agda.Builtin.Nat
open import Agda.Builtin.String

-- Non-dependent, strictly-positive pairs
record _×_ (@++ A B : Set) : Set where
  constructor _,_
  field fst : A
  field snd : B
open _×_

-- Names of syntax constructors of an ML-like language
data SynHead : Set where
  piTy natTy : SynHead
  lam app var : SynHead
  plus times natLit : SynHead

-- This passes:
synArgs : (@++ exp : Set)  SynHead  Set
synArgs exp piTy    = exp × exp
synArgs exp natTy   = ⊤
synArgs exp lam     = String × exp
synArgs exp var     = String
synArgs exp app     = exp × exp
synArgs exp natLit  = Nat
synArgs exp plus    = exp × exp
synArgs exp times   = exp × exp

-- This does *not* pass polarity checking
-- (I expected it to be equivalent):
module _ (@++ exp : Set) where
  synArgs' : SynHead  Set
  synArgs' piTy    = exp × exp
  synArgs' natTy   = ⊤
  synArgs' lam     = String × exp
  synArgs' var     = String
  synArgs' app     = exp × exp
  synArgs' natLit  = Nat
  synArgs' plus    = exp × exp
  synArgs' times   = exp × exp

The error message is

/home/lawrence/pre-release/east/src/bugreport.agda:35,22-25
Variable exp is bound with strictly positive polarity, so it cannot
be used here at a mixed position
when checking that the expression exp has type Set

My Agda version is the latest commit (4573e2d) of the un-merged PR #6385.

I'm reporting via a new issue to avoid blocking the merge (no existing code is broken).

This passes too:

postulate @++ exp : Set

synArgs'' : SynHead  Set
synArgs'' piTy    = exp × exp
synArgs'' natTy   = ⊤
synArgs'' lam     = String × exp
synArgs'' var     = String
synArgs'' app     = exp × exp
synArgs'' natLit  = Nat
synArgs'' plus    = exp × exp
synArgs'' times   = exp × exp

Does it pass if you mark the definition as @++? IIRC it's because all definitions are marked mixed by default, thus dividing the context by mixed, as alluded to in the doc. It's probably not fleshed out enough though.

Does it pass if you mark the definition as @++?

Yes, writing @++ synArgs' passes.

IIRC it's because all definitions are marked mixed by default, thus dividing the context by mixed, as alluded to in the doc. It's probably not fleshed out enough though.

Thanks for the explanation.

Here is a minimal example:

{-# OPTIONS --polarity #-}

p : (@++ A : Set)  Set
p A = q where
  @mixed q : Set
  q = A  -- Fails

module _ (@++ A : Set) where
  @mixed m : Set
  m = A    -- Fails

@mixed n : (@++ A : Set)  Set
n A = A   -- Passes

Actual behavior: Agda rejects m (like q)

Expected behavior: Agda accepts m (like n)

I read the modalities doc page, and it seems that this difference is due to #6583 being pure, not positional -- is this correct, @jpoiret?

And if so, what is the reason for choosing pure instead of positional?