Skip to content

Commit

Permalink
RulePattern: fix assignment of universe mvars in openRuleType
Browse files Browse the repository at this point in the history
The new Lean version seems to have broken this, possibly via changes to
`checkedAssign`.
  • Loading branch information
JLimperg committed Aug 6, 2024
1 parent 8b31db8 commit 2ca67d8
Show file tree
Hide file tree
Showing 2 changed files with 26 additions and 25 deletions.
6 changes: 3 additions & 3 deletions Aesop/RulePattern.lean
Original file line number Diff line number Diff line change
Expand Up @@ -145,9 +145,9 @@ def openRuleType (pat : RulePattern) (inst : RulePatternInstantiation)
for h : i in [:mvars.size] do
if let some inst ← pat.getInstantiation inst i then
let mvarId := mvars[i]'h.2 |>.mvarId!
-- We use `checkedAssign`, rather than `assign`, to make sure that
-- universe metavariables occurring in `mvars` are assigned.
ifmvarId.checkedAssign inst then
-- We use `isDefEq` to make sure that universe metavariables occurring in
-- the type of `mvarId` are assigned.
ifisDefEq (.mvar mvarId) inst then
assigned := assigned.insert mvarId
else
throwError "openRuleType: type-incorrect pattern instantiation: argument has type '{← mvarId.getType}' but pattern instantiation '{inst}' has type '{← inferType inst}'"
Expand Down
45 changes: 23 additions & 22 deletions AesopTest/RulePatternUniverseBug.lean
Original file line number Diff line number Diff line change
Expand Up @@ -8,7 +8,10 @@ Authors: Son Ho, Jannis Limperg
-- assigned during rule pattern matching. Thanks to Son Ho for reporting this
-- issue.

import Aesop
import Aesop.Frontend.Attribute
import Aesop.Frontend.Saturate

set_option aesop.check.all true

namespace List

Expand All @@ -28,24 +31,22 @@ def indexOpt (ls : List α) (i : Int) : Option α :=
| [] => none
| hd :: tl => if i = 0 then some hd else indexOpt tl (i - 1)

-- FIXME this is failing under v4.11.0-rc1

-- theorem indexOpt_bounds (ls : List α) (i : Int) :
-- ls.indexOpt i = none ↔ i < 0 ∨ ls.len ≤ i := by
-- match ls with
-- | [] => simp [indexOpt]; omega
-- | _ :: tl =>
-- have := indexOpt_bounds tl (i - 1)
-- if h : i = 0 then
-- simp [indexOpt, *]
-- saturate
-- omega
-- else
-- simp [indexOpt, len, *]
-- constructor <;> intro a <;> cases a
-- . left
-- saturate
-- omega
-- . right; omega
-- . left; omega
-- . right; omega
theorem indexOpt_bounds (ls : List α) (i : Int) :
ls.indexOpt i = none ↔ i < 0 ∨ ls.len ≤ i := by
match ls with
| [] => simp [indexOpt]; omega
| _ :: tl =>
have := indexOpt_bounds tl (i - 1)
if h : i = 0 then
simp [indexOpt, *]
saturate
omega
else
simp [indexOpt, len, *]
constructor <;> intro a <;> cases a
. left
saturate
omega
. right; omega
. left; omega
. right; omega

0 comments on commit 2ca67d8

Please sign in to comment.