Skip to content

Commit

Permalink
feat: add ac_nf and add test cases for BitVec
Browse files Browse the repository at this point in the history
ac_nf is a counterpart to ac_rfl, which normalizes bitvector expressions
with respect to associativity and commutativity.
  • Loading branch information
tobiasgrosser committed Sep 29, 2024
1 parent 4cd4bcc commit 70386a1
Show file tree
Hide file tree
Showing 3 changed files with 129 additions and 13 deletions.
4 changes: 4 additions & 0 deletions src/Init/Tactics.lean
Original file line number Diff line number Diff line change
Expand Up @@ -399,6 +399,10 @@ example (a b c d : Nat) : a + b + c + d = d + (b + c) + a := by ac_rfl
-/
syntax (name := acRfl) "ac_rfl" : tactic

/--
-/
syntax (name := acNf) "ac_nf" : tactic

/--
The `sorry` tactic closes the goal using `sorryAx`. This is intended for stubbing out incomplete
parts of a proof while still having a syntactically correct proof skeleton. Lean will give
Expand Down
41 changes: 28 additions & 13 deletions src/Lean/Meta/Tactic/AC/Main.lean
Original file line number Diff line number Diff line change
Expand Up @@ -140,19 +140,7 @@ where
| .op l r => mkApp2 preContext.op (convertTarget vars l) (convertTarget vars r)
| .var x => vars[x]!

def rewriteUnnormalized (mvarId : MVarId) : MetaM Unit := do
let simpCtx :=
{
simpTheorems := {}
congrTheorems := (← getSimpCongrTheorems)
config := Simp.neutralConfig
}
let tgt ← instantiateMVars (← mvarId.getType)
let (res, _) ← Simp.main tgt simpCtx (methods := { post })
let newGoal ← applySimpResultToTarget mvarId tgt res
newGoal.refl
where
post (e : Expr) : SimpM Simp.Step := do
def post (e : Expr) : SimpM Simp.Step := do
let ctx ← Simp.getContext
match e, ctx.parent? with
| bin op₁ l r, some (bin op₂ _ _) =>
Expand All @@ -171,10 +159,37 @@ where
| none => return Simp.Step.done { expr := e }
| e, _ => return Simp.Step.done { expr := e }

def rewriteUnnormalized (mvarId : MVarId) : MetaM Unit := do
let simpCtx :=
{
simpTheorems := {}
congrTheorems := (← getSimpCongrTheorems)
config := Simp.neutralConfig
}
let tgt ← instantiateMVars (← mvarId.getType)
let (res, _) ← Simp.main tgt simpCtx (methods := { post })
let newGoal ← applySimpResultToTarget mvarId tgt res
newGoal.refl

@[builtin_tactic acRfl] def acRflTactic : Lean.Elab.Tactic.Tactic := fun _ => do
let goal ← getMainGoal
goal.withContext <| rewriteUnnormalized goal

def canonicalizeUnnormalized (mvarId : MVarId) : TacticM Unit := do
let simpCtx :=
{
simpTheorems := {}
congrTheorems := (← getSimpCongrTheorems)
config := Simp.neutralConfig
}
let tgt ← instantiateMVars (← mvarId.getType)
let (res, _) ← Simp.main tgt simpCtx (methods := { post })
replaceMainGoal [← applySimpResultToTarget mvarId tgt res]

@[builtin_tactic acNf] def acNfTactic : Lean.Elab.Tactic.Tactic := fun _ => do
let goal ← getMainGoal
goal.withContext <| canonicalizeUnnormalized goal

builtin_initialize
registerTraceClass `Meta.AC

Expand Down
97 changes: 97 additions & 0 deletions tests/lean/run/ac.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,97 @@
/- arithmetic | commutativity -/

example (a b c d : BitVec w) :
a * b * c * d = d * c * b * a := by
ac_nf
rfl

example (a b c d : BitVec w) :
a * b * c * d = d * c * b * a := by
ac_rfl

example (a b c d : BitVec w) :
a + b + c + d = d + c + b + a := by
ac_nf
rfl

example (a b c d : BitVec w) :
a + b + c + d = d + c + b + a := by
ac_rfl

/- arithmetic | associativity -/

example (a b c d : BitVec w) :
a * (b * (c * d)) = ((a * b) * c) * d := by
ac_nf
rfl

example (a b c d : BitVec w) :
a * (b * (c * d)) = ((a * b) * c) * d := by
ac_rfl

example (a b c d : BitVec w) :
a + (b + (c + d)) = ((a + b) + c) + d := by
ac_nf
rfl

example (a b c d : BitVec w) :
a + (b + (c + d)) = ((a + b) + c) + d := by
ac_rfl

/- bitwise operations | commutativity -/

example (a b c d : BitVec w) :
a ^^^ b ^^^ c ^^^ d = d ^^^ c ^^^ b ^^^ a := by
ac_nf
rfl

example (a b c d : BitVec w) :
a ^^^ b ^^^ c ^^^ d = d ^^^ c ^^^ b ^^^ a := by
ac_rfl

example (a b c d : BitVec w) :
a &&& b &&& c &&& d = d &&& c &&& b &&& a := by
ac_nf
rfl

example (a b c d : BitVec w) :
a &&& b &&& c &&& d = d &&& c &&& b &&& a := by
ac_rfl

example (a b c d : BitVec w) :
a ||| b ||| c ||| d = d ||| c ||| b ||| a := by
ac_nf
rfl

example (a b c d : BitVec w) :
a ||| b ||| c ||| d = d ||| c ||| b ||| a := by
ac_rfl

/- bitwise operations | associativity -/

example (a b c d : BitVec w) :
a &&& (b &&& (c &&& d)) = ((a &&& b) &&& c) &&& d := by
ac_nf
rfl

example (a b c d : BitVec w) :
a &&& (b &&& (c &&& d)) = ((a &&& b) &&& c) &&& d := by
ac_rfl

example (a b c d : BitVec w) :
a ||| (b ||| (c ||| d)) = ((a ||| b) ||| c) ||| d := by
ac_nf
rfl

example (a b c d : BitVec w) :
a ||| (b ||| (c ||| d)) = ((a ||| b) ||| c) ||| d := by
ac_rfl

example (a b c d : BitVec w) :
a ^^^ (b ^^^ (c ^^^ d)) = ((a ^^^ b) ^^^ c) ^^^ d := by
ac_nf
rfl

example (a b c d : BitVec w) :
a ^^^ (b ^^^ (c ^^^ d)) = ((a ^^^ b) ^^^ c) ^^^ d := by
ac_rfl

0 comments on commit 70386a1

Please sign in to comment.