leanprover / lean4

Lean 4 programming language and theorem prover
https://lean-lang.org
Apache License 2.0
4.75k stars 428 forks source link

RFC: `@[flat]` annotation for names in the `extend` clause of a structure #2666

Open eric-wieser opened 1 year ago

eric-wieser commented 1 year ago

Proposal

What?

When working with a structure hierarchy like

structure A where a : Unit
structure B extends A where b : Unit
structure C extends A where c : Unit
structure D extends B, C where

the type of D.mk is asymmetric:

D.mk (toB : B) (c : Unit) : D

More precisely, the representation of D is to embed an entire B structure, but then only take the parts of the C structure that do not overlap with B, namely c : Unit.

I have no doubt that this behavior is useful in some circumstances, but in situations where a symmetric API is more important than performance, this behavior is annoying.

Let's imagine that the symmetric API we want is D.mk (toA : A) (b c : Unit). Right now, there are the following workarounds to achieve that

  1. Write the type of D manually without extends:

    structure A where a : Unit
    structure B extends A where b : Unit
    structure C extends A where c : Unit
    structure D extends A where
     b : Unit
     c : Unit
    
    -- user has to write this boilerplate
    def D.toB (d : D) : B := { toA := d.toA, b := d.b }
    def D.toC (d : D) : C := { toA := d.toA, c := d.c }

    This boilerplate can be rather a burden, epsecially since the naive := { d with } approach performs unwanted eta-expansion

  2. Use the hack from https://arxiv.org/abs/2306.00617 §4.2, which exploits the 'avoid overlaps' rule to our advantage, as

    structure DFlatHack where
    
    structure A where a : Unit
    structure B extends DFlaHack , A where b : Unit
    structure C extends DFlatHack, A where c : Unit
    structure D extends DFlatHack, B, C where

    which gives us D.mk (toDFlatHack : DFlatHack) (toA : A) (b c : Unit) : D, where the first argument is vacuous and can always be passed as ⟨⟩.

In Lean 3 we had set_option old_structure_cmd true; but this was a bad design because it did not allow parent structures to selectively be flattened (among other reasons).

The proposal is to remove the need for a hack here, such that the user can write

structure A where a : Unit
structure B extends A where b : Unit
structure C extends A where c : Unit
structure D extends @[flat] B, @[flat] C where

or any other reasonable syntax.

How?

The presence of flat simply means "skip the overlapping fields check, and behave as if the fields overlap"; that is, a check for the syntactic marker would be added to the conditions here:

https://github.com/leanprover/lean4/blob/57e23917b605d1f1e6b6ff70aa022065d55bb60c/src/Lean/Elab/Structure.lean#L423-L428

https://github.com/leanprover/lean4/blob/57e23917b605d1f1e6b6ff70aa022065d55bb60c/src/Lean/Elab/Structure.lean#L479-L483

Why?

In both cases, having a @[flat] attribute would make it vastly easier to perform performance and API experiments in mathlib to work out where nested structures help, and where they're a hindrance.

Community Feedback

Mario originally proposed this syntax here.

Impact

Add :+1: to issues you consider important. If others benefit from the changes in this proposal being added, please ask them to add :+1: to it.

mattrobball commented 1 year ago

Having concrete data on which patterns help where in mathlib would be very helpful.

eric-wieser commented 11 months ago

In both cases, having a @[flat] attribute would make it vastly easier to perform performance and API experiments in mathlib to work out where nested structures help, and where they're a hindrance.

In https://github.com/leanprover/lean4/pull/2940, I add a flat keyword (easier than new @[flat] notation), which allows us to investigate the effect of this in Mathlib. A mathlib toolchain that builds against this version can be found in https://github.com/leanprover-community/mathlib4/pull/8977.

eric-wieser commented 11 months ago

Some initial data: used on only the HomClass classes (a minimally invasive change), this provides a 2.5% wall clock boost for Mathlib.