Skip to content

Conversation

gabriellisboaconegero
Copy link

This pull request aims to continue and complete the work started in #2004

The first version already includes changes requested by @jamesmckinna, such as correct stdlib naming variables and redefinition of PropositionalNaperian.

@jamesmckinna
Copy link
Contributor

Thanks very much for continuing to work on this, and for responding to the earlier review comments.

I refactored the current version, as a way of trying to isolate the details of the contribtion; I'm happy to make local suggestions in a review, but I still think a global change is needed, as follows:

  • simplified the imports
  • lifted out the parametrisation, via an anonymous module, to expose the common elements between the private definition of FA, and that of Naperian itself
  • folded that definition into Naperian as a private manifest field
  • renamed the constructions so that each local module definition corresponding to a given record is given the same name (Agda's namespace control allows this, and, I think, is that rare case where a pun on names aids comprehension)
  • then I realised that the (extensional/observational) definition of equality in the setoid FS, in terms of change-of-base along index, precisely permits one of the 'axiom' fields to be expressed in terms of the other, so lifted that out as a manifest field.

Accordingly:

module _ {F : Set a  Set b} c (RF : RawFunctor F) where

-- RawNaperian contains just the functions, not the proofs

  record RawNaperian : Set (suc (a ⊔ c) ⊔ b) where
    field
      Log : Set c
      index : F A  (Log  A)
      tabulate : (Log  A)  F A

-- Full Naperian has the coherence conditions too.

  record Naperian (S : Setoid a ℓ) : Set (suc (a ⊔ c) ⊔ b ⊔ ℓ) where
    field
      rawNaperian : RawNaperian
    open RawNaperian rawNaperian public
    open module S = Setoid S
    private
      FS : Setoid b (c ⊔ ℓ)
      FS = record
        { _≈_ = λ (fx fy : F Carrier)   (l : Log)  index fx l ≈ index fy l
        ; isEquivalence = record
          { refl = λ _  refl
          ; sym = λ eq l  sym (eq l)
          ; trans = λ i≈j j≈k l  trans (i≈j l) (j≈k l)
          }
        }
      module FS = Setoid FS
    field
      -- tabulate-index : (fx : F Carrier)  tabulate (index fx) FS.≈ fx
      index-tabulate : (f : Log  Carrier) (l : Log)  index (tabulate f) l ≈ f l
  
    tabulate-index : (fx : F Carrier)  tabulate (index fx) FS.≈ fx
    tabulate-index = index-tabulate ∘ index

  PropositionalNaperian : Set (suc (a ⊔ c) ⊔ b)
  PropositionalNaperian =  A  Naperian (≡.setoid A)

@jamesmckinna
Copy link
Contributor

jamesmckinna commented Aug 27, 2025

Now, after such (I think: semantics-preserving!) rearrangement, we can observe that the parametrisation on S : Setoid is still external to the definition of Naperian, whereas the parametrisation on A : Set a in the definition of RawNaperian is internal, ie., at least if I understand things correctly, that index and tabulate are insufficiently generalised: their types should, I think, depend on the underlying Carrier of a given Setoid argument, rather than being free-standing fields of a RawNaperian which are then applied to such a Carrier...

As things stand, Naperian doesn't define 'functorial' behaviour; rather it collects properties that for a given S : Setoid, says that some underlying RawNaperian raw functor is extensional on F S, rather than specifying such behaviour parametrically. In formalistic terms, it also permits a given Naperian to choose a different underlying RawNaperian for each choice of S, which is surely not what we want for a given Functor?

Hope that the suggested refactoring make sit easier to see a possible way forward, but as I've critiqued the previous design in this style, a shout-out to @JacquesCarette for comment on this latest iteration.

@jamesmckinna
Copy link
Contributor

Re: CHANGELOG
This is a horrible mess caused by this PR basing off the old, pre-v2.3 version, so perhaps the best thing to do is to push a 'fresh' copy of CHANGELOG on master to his branch, and then add the small patch needed to record the addition of this one new module? But others with superior git-fu (mine is still very rudimentary) may have a better suggestion!

Copy link
Contributor

@jamesmckinna jamesmckinna left a comment

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Lots that can be (further) simplified, but I still think that the parametrisation on Setoid needs fixing!

module Effect.Functor.Naperian where

open import Effect.Functor using (RawFunctor)
open import Function.Bundles using (_⟶ₛ_; _⟨$⟩_; Func)
Copy link
Contributor

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Suggested change
open import Function.Bundles using (_⟶ₛ_; _⟨$⟩_; Func)
open import Function.Base using (_∘_)

open import Function.Bundles using (_⟶ₛ_; _⟨$⟩_; Func)
open import Level using (Level; suc; _⊔_)
open import Relation.Binary.Bundles using (Setoid)
open import Relation.Binary.PropositionalEquality.Core using (_≡_)
Copy link
Contributor

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

The relation itself is now never used, only its property of admitting a Setoid bundle, so this can be deleted

Suggested change
open import Relation.Binary.PropositionalEquality.Core using (_≡_)

open import Level using (Level; suc; _⊔_)
open import Relation.Binary.Bundles using (Setoid)
open import Relation.Binary.PropositionalEquality.Core using (_≡_)
open import Relation.Binary.PropositionalEquality.Properties as ≡
Copy link
Contributor

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Can simplify this to:

Suggested change
open import Relation.Binary.PropositionalEquality.Properties as ≡
import Relation.Binary.PropositionalEquality.Properties as ≡ using (setoid)


private
variable
a b c e f : Level
Copy link
Contributor

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Standardise, based on parametrisation on S : Setoid a ℓ, itself a lexical convention:

Suggested change
a b c e f : Level
a b c : Level

Comment on lines 45 to 52
-- module Propositional where
-- record Naperian {F : Set a → Set b} c (RF : RawFunctor F) : Set (suc (a ⊔ c) ⊔ b) where
-- field
-- rawNaperian : RawNaperian c RF
-- open RawNaperian rawNaperian public
-- field
-- tabulate-index : (fa : F A) → tabulate (index fa) ≡ fa
-- index-tabulate : (f : Log → A) → ((l : Log) → index (tabulate f) l ≡ f l)
Copy link
Contributor

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Delete!

Comment on lines 54 to 79
module _ {F : Set a → Set b} c (RF : RawFunctor F) where
private
FA : (S : Setoid a e) → (rn : RawNaperian c RF) → Setoid b (c ⊔ e)
FA S rn = record
{ _≈_ = λ (fx fy : F Carrier) → (l : Log) → index fx l ≈ index fy l
; isEquivalence = record
{ refl = λ _ → refl
; sym = λ eq l → sym (eq l)
; trans = λ i≈j j≈k l → trans (i≈j l) (j≈k l)
}
}
where
open Setoid S
open RawNaperian rn

record Naperian (S : Setoid a e) : Set (suc a ⊔ b ⊔ suc c ⊔ e) where
field
rawNaperian : RawNaperian c RF
open RawNaperian rawNaperian public
private
module FS = Setoid (FA S rawNaperian)
module A = Setoid S
field
tabulate-index : (fx : F A.Carrier) → tabulate (index fx) FS.≈ fx
index-tabulate : (f : Log → A.Carrier) → ((l : Log) → index (tabulate f) l A.≈ f l)

Copy link
Contributor

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

See the revised version in my extended comment.

index-tabulate : (f : Log → A.Carrier) → ((l : Log) → index (tabulate f) l A.≈ f l)

PropositionalNaperian : Set (suc (a ⊔ c) ⊔ b)
PropositionalNaperian = ∀ {A} → Naperian (≡.setoid A)
Copy link
Contributor

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

On the current design wrt parametrisation, I think that A should be explicit:

Suggested change
PropositionalNaperian = {A} Naperian (≡.setoid A)
PropositionalNaperian = A Naperian (≡.setoid A)

@JacquesCarette
Copy link
Contributor

Yes, it would make sense to 'restart' the work on CHANGELOG from a recent version, as the merge is otherwise going to be hell.

@JacquesCarette
Copy link
Contributor

Thanks for the review @jamesmckinna . All the explicit suggestions should go in - please do so @gabriellisboaconegero .

Comment on lines +64 to +65
natural-tabulate : (f : Carrier → Carrier) (k : Log → Carrier) → (tabulate (f ∘ k)) FS.≈ (f <$> (tabulate k))
natural-index : (f : Carrier → Carrier) (as : F Carrier) (l : Log) → (index (f <$> as) l) ≈ f (index as l)
Copy link
Contributor

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Nice!
Can natural-tabulate also be eliminated by defining it in terms of index-tabulate, natural-index, and tabulate-index?

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

I don't think it can without assuming congruence on the setoid. I was trying to do so, but I was stuck and needed to prove the function congruence on the setoid S. I could figure out that

natural-index f (tabulate k) l -> index (f <$> tabulate k) l ≈ f (index (tabulate k) l)
index-tabulate k l                    -> index (tabulate k) l ≈ k l
index-tabulate (f ∘ k) l            -> index (tabulate (λ x  f (k x))) l ≈ f (k l)

If I assume congruence, then I could prove tabulate-index.

I don't think we can find a proof for tabulate-index without congruence, but I am not sure. I base my argument on the fact that tabulate-∘ needs cong to be proved.

Sign up for free to join this conversation on GitHub. Already have an account? Sign in to comment
Labels
None yet
Projects
None yet
Development

Successfully merging this pull request may close these issues.

4 participants