summaryrefslogtreecommitdiff
path: root/backends/lean/Base
diff options
context:
space:
mode:
Diffstat (limited to 'backends/lean/Base')
-rw-r--r--backends/lean/Base/Arith/Int.lean6
-rw-r--r--backends/lean/Base/Arith/Scalar.lean57
-rw-r--r--backends/lean/Base/Primitives/Scalar.lean7
-rw-r--r--backends/lean/Base/Progress/Progress.lean16
-rw-r--r--backends/lean/Base/Utils.lean77
5 files changed, 118 insertions, 45 deletions
diff --git a/backends/lean/Base/Arith/Int.lean b/backends/lean/Base/Arith/Int.lean
index 4a3db5f8..59cdca25 100644
--- a/backends/lean/Base/Arith/Int.lean
+++ b/backends/lean/Base/Arith/Int.lean
@@ -180,7 +180,7 @@ def introInstances (declToUnfold : Name) (lookup : Expr → MetaM (Option Expr))
-- Add a declaration
let nval ← Utils.addDeclTac name e type (asLet := false)
-- Simplify to unfold the declaration to unfold (i.e., the projector)
- Utils.simpAt true [declToUnfold] [] [] (Location.targets #[mkIdent name] false)
+ Utils.simpAt true {} [declToUnfold] [] [] (Location.targets #[mkIdent name] false)
-- Return the new value
pure nval
@@ -214,7 +214,7 @@ def intTacPreprocess (extraPreprocess : Tactic.TacticM Unit) : Tactic.TacticM U
extraPreprocess
-- Reduce all the terms in the goal - note that the extra preprocessing step
-- might have proven the goal, hence the `Tactic.allGoals`
- Tactic.allGoals do tryTac (dsimpAt false [] [] [] Tactic.Location.wildcard)
+ Tactic.allGoals do tryTac (dsimpAt false {} [] [] [] Tactic.Location.wildcard)
elab "int_tac_preprocess" : tactic =>
intTacPreprocess (do pure ())
@@ -231,7 +231,7 @@ def intTac (tacName : String) (splitGoalConjs : Bool) (extraPreprocess : Tactic
-- the goal. I think before leads to a smaller proof term?
Tactic.allGoals (intTacPreprocess extraPreprocess)
-- More preprocessing
- Tactic.allGoals (Utils.tryTac (Utils.simpAt true [] [``nat_zero_eq_int_zero] [] .wildcard))
+ Tactic.allGoals (Utils.tryTac (Utils.simpAt true {} [] [``nat_zero_eq_int_zero] [] .wildcard))
-- Split the conjunctions in the goal
if splitGoalConjs then Tactic.allGoals (Utils.repeatTac Utils.splitConjTarget)
-- Call linarith
diff --git a/backends/lean/Base/Arith/Scalar.lean b/backends/lean/Base/Arith/Scalar.lean
index 86b2e216..8793713b 100644
--- a/backends/lean/Base/Arith/Scalar.lean
+++ b/backends/lean/Base/Arith/Scalar.lean
@@ -8,30 +8,29 @@ open Lean Lean.Elab Lean.Meta
open Primitives
def scalarTacExtraPreprocess : Tactic.TacticM Unit := do
- Tactic.withMainContext do
- -- Inroduce the bounds for the isize/usize types
- let add (e : Expr) : Tactic.TacticM Unit := do
- let ty ← inferType e
- let _ ← Utils.addDeclTac (← Utils.mkFreshAnonPropUserName) e ty (asLet := false)
- add (← mkAppM ``Scalar.cMin_bound #[.const ``ScalarTy.Isize []])
- add (← mkAppM ``Scalar.cMax_bound #[.const ``ScalarTy.Usize []])
- add (← mkAppM ``Scalar.cMax_bound #[.const ``ScalarTy.Isize []])
- -- Reveal the concrete bounds, simplify calls to [ofInt]
- Utils.simpAt true
- -- Unfoldings
- [``Scalar.min, ``Scalar.max, ``Scalar.cMin, ``Scalar.cMax,
- ``I8.min, ``I16.min, ``I32.min, ``I64.min, ``I128.min,
- ``I8.max, ``I16.max, ``I32.max, ``I64.max, ``I128.max,
- ``U8.min, ``U16.min, ``U32.min, ``U64.min, ``U128.min,
- ``U8.max, ``U16.max, ``U32.max, ``U64.max, ``U128.max,
- ``Usize.min
- ]
- -- Simp lemmas
- [``Scalar.ofInt_val_eq, ``Scalar.neq_to_neq_val,
- ``Scalar.lt_equiv, ``Scalar.le_equiv, ``Scalar.eq_equiv]
- -- Hypotheses
- [] .wildcard
-
+ Tactic.withMainContext do
+ -- Inroduce the bounds for the isize/usize types
+ let add (e : Expr) : Tactic.TacticM Unit := do
+ let ty ← inferType e
+ let _ ← Utils.addDeclTac (← Utils.mkFreshAnonPropUserName) e ty (asLet := false)
+ add (← mkAppM ``Scalar.cMin_bound #[.const ``ScalarTy.Isize []])
+ add (← mkAppM ``Scalar.cMax_bound #[.const ``ScalarTy.Usize []])
+ add (← mkAppM ``Scalar.cMax_bound #[.const ``ScalarTy.Isize []])
+ -- Reveal the concrete bounds, simplify calls to [ofInt]
+ Utils.simpAt true {}
+ -- Unfoldings
+ [``Scalar.min, ``Scalar.max, ``Scalar.cMin, ``Scalar.cMax,
+ ``I8.min, ``I16.min, ``I32.min, ``I64.min, ``I128.min,
+ ``I8.max, ``I16.max, ``I32.max, ``I64.max, ``I128.max,
+ ``U8.min, ``U16.min, ``U32.min, ``U64.min, ``U128.min,
+ ``U8.max, ``U16.max, ``U32.max, ``U64.max, ``U128.max,
+ ``Usize.min
+ ]
+ -- Simp lemmas
+ [``Scalar.ofInt_val_eq, ``Scalar.neq_to_neq_val,
+ ``Scalar.lt_equiv, ``Scalar.le_equiv, ``Scalar.eq_equiv]
+ -- Hypotheses
+ [] .wildcard
elab "scalar_tac_preprocess" : tactic =>
intTacPreprocess scalarTacExtraPreprocess
@@ -81,4 +80,14 @@ example (x : Int) (h0 : 0 ≤ x) (h1 : x ≤ U32.max) :
example (x : U32) (h0 : ¬ x = U32.ofInt 0) : 0 < x.val := by
scalar_tac
+/- See this: https://aeneas-verif.zulipchat.com/#narrow/stream/349819-general/topic/U64.20trouble/near/444049757
+
+ We solved it by removing the instance `OfNat` for `Scalar`.
+ Note however that we could also solve it with a simplification lemma.
+ However, after testing, we noticed we could only apply such a lemma with
+ the rewriting tactic (not the simplifier), probably because of the use
+ of typeclasses. -/
+example {u: U64} (h1: (u : Int) < 2): (u : Int) = 0 ∨ (u : Int) = 1 := by
+ scalar_tac
+
end Arith
diff --git a/backends/lean/Base/Primitives/Scalar.lean b/backends/lean/Base/Primitives/Scalar.lean
index 8fb067e1..157ade2c 100644
--- a/backends/lean/Base/Primitives/Scalar.lean
+++ b/backends/lean/Base/Primitives/Scalar.lean
@@ -351,10 +351,17 @@ instance [Decide (Scalar.cMin ty ≤ v ∧ v ≤ Scalar.cMax ty)] : InBounds ty
@[simp] abbrev Scalar.check_bounds (ty : ScalarTy) (x : Int) : Bool :=
(Scalar.cMin ty ≤ x || Scalar.min ty ≤ x) ∧ (x ≤ Scalar.cMax ty || x ≤ Scalar.max ty)
+/- Discussion:
+ This coercion can be slightly annoying at times, because if we write
+ something like `u = 3` (where `u` is, for instance, as `U32`), then instead of
+ coercing `u` to `Int`, Lean will lift `3` to `U32`).
+ For now we deactivate it.
+
-- TODO(raitobezarius): the inbounds constraint is a bit ugly as we can pretty trivially
-- discharge the lhs on ≥ 0.
instance {ty: ScalarTy} [InBounds ty (Int.ofNat n)]: OfNat (Scalar ty) (n: ℕ) where
ofNat := Scalar.ofInt n
+-/
theorem Scalar.check_bounds_imp_in_bounds {ty : ScalarTy} {x : Int}
(h: Scalar.check_bounds ty x) :
diff --git a/backends/lean/Base/Progress/Progress.lean b/backends/lean/Base/Progress/Progress.lean
index f2a56e50..03d464d7 100644
--- a/backends/lean/Base/Progress/Progress.lean
+++ b/backends/lean/Base/Progress/Progress.lean
@@ -58,17 +58,13 @@ def progressWith (fExpr : Expr) (th : TheoremOrLocal)
We also make sure that all the meta variables which appear in the
function arguments have been instantiated
-/
- let env ← getEnv
let thTy ← do
match th with
| .Theorem thName =>
- let thDecl := env.constants.find! thName
- -- We have to introduce fresh meta-variables for the universes already
- let ul : List (Name × Level) ←
- thDecl.levelParams.mapM (λ x => do pure (x, ← mkFreshLevelMVar))
- let ulMap : HashMap Name Level := HashMap.ofList ul
- let thTy := thDecl.type.instantiateLevelParamsCore (λ x => ulMap.find! x)
- pure thTy
+ -- Lookup the theorem and introduce fresh meta-variables for the universes
+ let th ← mkConstWithFreshMVarLevels thName
+ -- Retrieve the type
+ inferType th
| .Local asmDecl => pure asmDecl.type
trace[Progress] "Looked up theorem/assumption type: {thTy}"
-- TODO: the tactic fails if we uncomment withNewMCtxDepth
@@ -135,7 +131,7 @@ def progressWith (fExpr : Expr) (th : TheoremOrLocal)
Tactic.focus do
let _ ←
tryTac
- (simpAt true []
+ (simpAt true {} []
[``Primitives.bind_tc_ok, ``Primitives.bind_tc_fail, ``Primitives.bind_tc_div]
[hEq.fvarId!] (.targets #[] true))
-- It may happen that at this point the goal is already solved (though this is rare)
@@ -144,7 +140,7 @@ def progressWith (fExpr : Expr) (th : TheoremOrLocal)
else
trace[Progress] "goal after applying the eq and simplifying the binds: {← getMainGoal}"
-- TODO: remove this (some types get unfolded too much: we "fold" them back)
- let _ ← tryTac (simpAt true [] scalar_eqs [] .wildcard_dep)
+ let _ ← tryTac (simpAt true {} [] scalar_eqs [] .wildcard_dep)
trace[Progress] "goal after folding back scalar types: {← getMainGoal}"
-- Clear the equality, unless the user requests not to do so
let mgoal ← do
diff --git a/backends/lean/Base/Utils.lean b/backends/lean/Base/Utils.lean
index 7ae5a832..4be46400 100644
--- a/backends/lean/Base/Utils.lean
+++ b/backends/lean/Base/Utils.lean
@@ -664,7 +664,7 @@ example (h : ∃ x y z, x + y + z ≥ 0) : ∃ x, x ≥ 0 := by
Something very annoying is that there is no function which allows to
initialize a simp context without doing an elaboration - as a consequence
we write our own here. -/
-def mkSimpCtx (simpOnly : Bool) (declsToUnfold : List Name) (thms : List Name) (hypsToUse : List FVarId) :
+def mkSimpCtx (simpOnly : Bool) (config : Simp.Config) (declsToUnfold : List Name) (thms : List Name) (hypsToUse : List FVarId) :
Tactic.TacticM Simp.Context := do
-- Initialize either with the builtin simp theorems or with all the simp theorems
let simpThms ←
@@ -693,7 +693,7 @@ def mkSimpCtx (simpOnly : Bool) (declsToUnfold : List Name) (thms : List Name) (
throwError "Not a proposition: {thmName}"
) simpThms
let congrTheorems ← getSimpCongrTheorems
- pure { simpTheorems := #[simpThms], congrTheorems }
+ pure { config, simpTheorems := #[simpThms], congrTheorems }
inductive Location where
/-- Apply the tactic everywhere. Same as `Tactic.Location.wildcard` -/
@@ -731,29 +731,90 @@ where
return usedSimps
/- Call the simp tactic. -/
-def simpAt (simpOnly : Bool) (declsToUnfold : List Name) (thms : List Name) (hypsToUse : List FVarId)
+def simpAt (simpOnly : Bool) (config : Simp.Config) (declsToUnfold : List Name) (thms : List Name) (hypsToUse : List FVarId)
(loc : Location) :
Tactic.TacticM Unit := do
-- Initialize the simp context
- let ctx ← mkSimpCtx simpOnly declsToUnfold thms hypsToUse
+ let ctx ← mkSimpCtx simpOnly config declsToUnfold thms hypsToUse
-- Apply the simplifier
let _ ← customSimpLocation ctx (discharge? := .none) loc
/- Call the dsimp tactic. -/
-def dsimpAt (simpOnly : Bool) (declsToUnfold : List Name) (thms : List Name) (hypsToUse : List FVarId)
+def dsimpAt (simpOnly : Bool) (config : Simp.Config) (declsToUnfold : List Name) (thms : List Name) (hypsToUse : List FVarId)
(loc : Tactic.Location) :
Tactic.TacticM Unit := do
-- Initialize the simp context
- let ctx ← mkSimpCtx simpOnly declsToUnfold thms hypsToUse
+ let ctx ← mkSimpCtx simpOnly config declsToUnfold thms hypsToUse
-- Apply the simplifier
dsimpLocation ctx loc
-- Call the simpAll tactic
-def simpAll (declsToUnfold : List Name) (thms : List Name) (hypsToUse : List FVarId) :
+def simpAll (config : Simp.Config) (declsToUnfold : List Name) (thms : List Name) (hypsToUse : List FVarId) :
Tactic.TacticM Unit := do
-- Initialize the simp context
- let ctx ← mkSimpCtx false declsToUnfold thms hypsToUse
+ let ctx ← mkSimpCtx false config declsToUnfold thms hypsToUse
-- Apply the simplifier
let _ ← Lean.Meta.simpAll (← getMainGoal) ctx
+/- Adapted from Elab.Tactic.Rewrite -/
+def rewriteTarget (eqThm : Expr) (symm : Bool) (config : Rewrite.Config := {}) : TacticM Unit := do
+ Term.withSynthesize <| withMainContext do
+ let r ← (← getMainGoal).rewrite (← getMainTarget) eqThm symm (config := config)
+ let mvarId' ← (← getMainGoal).replaceTargetEq r.eNew r.eqProof
+ replaceMainGoal (mvarId' :: r.mvarIds)
+
+/- Adapted from Elab.Tactic.Rewrite -/
+def rewriteLocalDecl (eqThm : Expr) (symm : Bool) (fvarId : FVarId) (config : Rewrite.Config := {}) :
+ TacticM Unit := withMainContext do
+ -- Note: we cannot execute `replaceLocalDecl` inside `Term.withSynthesize`.
+ -- See issues #2711 and #2727.
+ let rwResult ← Term.withSynthesize <| withMainContext do
+ let localDecl ← fvarId.getDecl
+ (← getMainGoal).rewrite localDecl.type eqThm symm (config := config)
+ let replaceResult ← (← getMainGoal).replaceLocalDecl fvarId rwResult.eNew rwResult.eqProof
+ replaceMainGoal (replaceResult.mvarId :: rwResult.mvarIds)
+
+/- Adapted from Elab.Tactic.Rewrite -/
+def rewriteWithThms
+ (thms : List (Bool × Expr))
+ (rewrite : (symm : Bool) → (thm : Expr) → TacticM Unit)
+ : TacticM Unit := do
+ let rec go thms :=
+ match thms with
+ | [] => throwError "Failed to rewrite with any theorem"
+ | (symm, eqThm)::thms =>
+ rewrite symm eqThm <|> go thms
+ go thms
+
+/- Adapted from Elab.Tactic.Rewrite -/
+def evalRewriteSeqAux (cfg : Rewrite.Config) (thms : List (Bool × Expr)) (loc : Tactic.Location) : TacticM Unit :=
+ rewriteWithThms thms fun symm term => do
+ withLocation loc
+ (rewriteLocalDecl term symm · cfg)
+ (rewriteTarget term symm cfg)
+ (throwTacticEx `rewrite · "did not find instance of the pattern in the current goal")
+
+/-- `rpt`: if `true`, repeatedly rewrite -/
+def rewriteAt (cfg : Rewrite.Config) (rpt : Bool)
+ (thms : List (Bool × Name)) (loc : Tactic.Location) : TacticM Unit := do
+ -- Lookup the theorems
+ let lookupThm (x : Bool × Name) : TacticM (List (Bool × Expr)) := do
+ let thName := x.snd
+ let lookupOne (thName : Name) : TacticM (Bool × Expr) := do
+ -- Lookup the theorem and introduce fresh meta-variables for the universes
+ let th ← mkConstWithFreshMVarLevels thName
+ pure (x.fst, th)
+ match ← getEqnsFor? thName (nonRec := true) with
+ | some eqThms => do
+ eqThms.data.mapM lookupOne
+ | none => do
+ pure [← lookupOne thName]
+ let thms ← List.mapM lookupThm thms
+ let thms := thms.flatten
+ -- Rewrite
+ if rpt then
+ Utils.repeatTac (evalRewriteSeqAux cfg thms loc)
+ else
+ evalRewriteSeqAux cfg thms loc
+
end Utils