summaryrefslogtreecommitdiff
path: root/backends/lean/Base/Arith
diff options
context:
space:
mode:
Diffstat (limited to 'backends/lean/Base/Arith')
-rw-r--r--backends/lean/Base/Arith/Base.lean12
-rw-r--r--backends/lean/Base/Arith/Int.lean15
-rw-r--r--backends/lean/Base/Arith/Scalar.lean13
3 files changed, 37 insertions, 3 deletions
diff --git a/backends/lean/Base/Arith/Base.lean b/backends/lean/Base/Arith/Base.lean
index 9c11ed45..8ada4171 100644
--- a/backends/lean/Base/Arith/Base.lean
+++ b/backends/lean/Base/Arith/Base.lean
@@ -57,4 +57,16 @@ theorem int_pos_ind (p : Int → Prop) :
-- TODO: there is probably something more general to do
theorem nat_zero_eq_int_zero : (0 : Nat) = (0 : Int) := by simp
+-- This is mostly used in termination proofs
+theorem to_int_to_nat_lt (x y : ℤ) (h0 : 0 ≤ x) (h1 : x < y) :
+ ↑(x.toNat) < y := by
+ simp [*]
+
+-- This is mostly used in termination proofs
+theorem to_int_sub_to_nat_lt (x y : ℤ) (x' : ℕ)
+ (h0 : ↑x' ≤ x) (h1 : x - ↑x' < y) :
+ ↑(x.toNat - x') < y := by
+ have : 0 ≤ x := by linarith
+ simp [Int.toNat_sub_of_le, *]
+
end Arith
diff --git a/backends/lean/Base/Arith/Int.lean b/backends/lean/Base/Arith/Int.lean
index 3359ecdb..a57f8bb1 100644
--- a/backends/lean/Base/Arith/Int.lean
+++ b/backends/lean/Base/Arith/Int.lean
@@ -162,7 +162,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 [declToUnfold] [] [] (Tactic.Location.targets #[mkIdent name] false)
+ Utils.simpAt true [declToUnfold] [] [] (Location.targets #[mkIdent name] false)
-- Return the new value
pure nval
@@ -240,7 +240,7 @@ def intTac (splitGoalConjs : Bool) (extraPreprocess : Tactic.TacticM Unit) : Ta
-- the goal. I think before leads to a smaller proof term?
Tactic.allGoals (intTacPreprocess extraPreprocess)
-- More preprocessing
- Tactic.allGoals (Utils.tryTac (Utils.simpAt [] [``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
@@ -270,6 +270,17 @@ elab "int_tac" args:(" split_goal"?): tactic =>
let split := args.raw.getArgs.size > 0
intTac split (do pure ())
+-- For termination proofs
+syntax "int_decr_tac" : tactic
+macro_rules
+ | `(tactic| int_decr_tac) =>
+ `(tactic|
+ simp_wf;
+ -- TODO: don't use a macro (namespace problems)
+ (first | apply Arith.to_int_to_nat_lt
+ | apply Arith.to_int_sub_to_nat_lt) <;>
+ simp_all <;> int_tac)
+
example (x : Int) (h0: 0 ≤ x) (h1: x ≠ 0) : 0 < x := by
int_tac_preprocess
linarith
diff --git a/backends/lean/Base/Arith/Scalar.lean b/backends/lean/Base/Arith/Scalar.lean
index 47751c8a..2342cce6 100644
--- a/backends/lean/Base/Arith/Scalar.lean
+++ b/backends/lean/Base/Arith/Scalar.lean
@@ -17,7 +17,7 @@ def scalarTacExtraPreprocess : Tactic.TacticM Unit := do
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 [``Scalar.min, ``Scalar.max, ``Scalar.cMin, ``Scalar.cMax,
+ Utils.simpAt true [``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,
@@ -36,6 +36,17 @@ def scalarTac (splitGoalConjs : Bool) : Tactic.TacticM Unit := do
elab "scalar_tac" : tactic =>
scalarTac false
+-- For termination proofs
+syntax "scalar_decr_tac" : tactic
+macro_rules
+ | `(tactic| scalar_decr_tac) =>
+ `(tactic|
+ simp_wf;
+ -- TODO: don't use a macro (namespace problems)
+ (first | apply Arith.to_int_to_nat_lt
+ | apply Arith.to_int_sub_to_nat_lt) <;>
+ simp_all <;> scalar_tac)
+
instance (ty : ScalarTy) : HasIntProp (Scalar ty) where
-- prop_ty is inferred
prop := λ x => And.intro x.hmin x.hmax