Skip to content

Commit

Permalink
feat: add autoPromoteIndices option
Browse files Browse the repository at this point in the history
  • Loading branch information
arthur-adjedj committed Mar 4, 2024
1 parent 49f41a6 commit 8cdd365
Show file tree
Hide file tree
Showing 2 changed files with 35 additions and 1 deletion.
9 changes: 8 additions & 1 deletion src/Lean/Elab/Inductive.lean
Original file line number Diff line number Diff line change
Expand Up @@ -25,6 +25,11 @@ open Meta
builtin_initialize
registerTraceClass `Elab.inductive

register_builtin_option autoPromoteIndices : Bool := {
defValue := true
descr := "Promote indices to parameters in inductive types whenever possible."
}

def checkValidInductiveModifier [Monad m] [MonadError m] (modifiers : Modifiers) : m Unit := do
if modifiers.isNoncomputable then
throwError "invalid use of 'noncomputable' in inductive declaration"
Expand Down Expand Up @@ -714,10 +719,12 @@ private def isDomainDefEq (arrowType : Expr) (type : Expr) : MetaM Bool := do
Convert fixed indices to parameters.
-/
private partial def fixedIndicesToParams (numParams : Nat) (indTypes : Array InductiveType) (indFVars : Array Expr) : MetaM Nat := do
if !autoPromoteIndices.get (← getOptions) then
return numParams
let masks ← indTypes.mapM (computeFixedIndexBitMask numParams · indFVars)
trace[Elab.inductive] "masks: {masks}"
if masks.all fun mask => !mask.contains true then
return numParams
trace[Elab.inductive] "masks: {masks}"
-- We process just a non-fixed prefix of the indices for now. Reason: we don't want to change the order.
-- TODO: extend it in the future. For example, it should be reasonable to change
-- the order of indices generated by the auto implicit feature.
Expand Down
27 changes: 27 additions & 0 deletions tests/lean/run/3458_1.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,27 @@
/- Ensure that turning off `autoPromoteIndices` stops indices
from being promoted to parameters in inductive types.
Because free variables aren't allowed in parameters of nested appearances,
it is sometimes desirable for fixed indices to not be promoted to parameters.
See `IInterp`.
-/

set_option autoImplicit false
universe u v w

structure ISignature (I : Type u) : Type (max u v + 1) where
symbols : I → Type v
indices : {i : I} → symbols i → List I

inductive All {I : Type u} (P : I → Type v) : List I → Type (max u v) where
| nil : All P []
| cons {x xs}: P x → All P xs → All P (x :: xs)

set_option autoPromoteIndices false

inductive IInterp {I} (ζ : ISignature.{u,v} I) (P : I → Type w) : I → Type (max u v w) where
| mk :{i : I} → (s : ζ.symbols i) → All P (ζ.indices s) → IInterp ζ P i

notation:max "I⟦ " ζ "" => IInterp ζ

inductive {I : Type u} (ζ : ISignature.{u,v} I) : I → Type (max u v) where
| mk : (i : I) → I⟦ ζ ⟧ (Iμ ζ) i → Iμ ζ i

0 comments on commit 8cdd365

Please sign in to comment.