diff options
| author | Son Ho | 2024-06-17 06:16:43 +0200 | 
|---|---|---|
| committer | Son Ho | 2024-06-17 06:16:43 +0200 | 
| commit | e57e6f08e5cc34bf4e9237650f5ecbab440b9ea2 (patch) | |
| tree | 1e48b2d23719d72f39282213a1806591cc35c3b8 /backends/lean/Base/Arith | |
| parent | f3b22b5cca9bc1154f55a81c9a82dc491074067d (diff) | |
| parent | 85098d7caf5e3196c2e8f92411efd2814bfed1ea (diff) | |
Merge branch 'son/update-lean' into has-int-pred
Diffstat (limited to '')
| -rw-r--r-- | backends/lean/Base/Arith/Base.lean | 11 | ||||
| -rw-r--r-- | backends/lean/Base/Arith/Int.lean | 10 | ||||
| -rw-r--r-- | backends/lean/Base/Arith/Scalar.lean | 59 | 
3 files changed, 45 insertions, 35 deletions
diff --git a/backends/lean/Base/Arith/Base.lean b/backends/lean/Base/Arith/Base.lean index 8ada4171..fb6b12e5 100644 --- a/backends/lean/Base/Arith/Base.lean +++ b/backends/lean/Base/Arith/Base.lean @@ -1,6 +1,5 @@  import Lean -import Std.Data.Int.Lemmas -import Mathlib.Tactic.Linarith +import Mathlib.Tactic.Linarith -- Introduces a lot of useful lemmas  namespace Arith @@ -21,12 +20,12 @@ theorem ne_is_lt_or_gt {x y : Int} (hne : x ≠ y) : x < y ∨ x > y := by    have hne : x - y ≠ 0 := by      simp      intro h -    have: x = y := by linarith +    have: x = y := by omega      simp_all    have h := ne_zero_is_lt_or_gt hne    match h with -  | .inl _ => left; linarith -  | .inr _ => right; linarith +  | .inl _ => left; omega +  | .inr _ => right; omega  -- TODO: move?  theorem add_one_le_iff_le_ne (n m : Nat) (h1 : m ≤ n) (h2 : m ≠ n) : m + 1 ≤ n := by @@ -66,7 +65,7 @@ theorem to_int_to_nat_lt (x y : ℤ) (h0 : 0 ≤ x) (h1 : x < y) :  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 +  have : 0 ≤ x := by omega    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 a1cb9da3..ab6dd4ab 100644 --- a/backends/lean/Base/Arith/Int.lean +++ b/backends/lean/Base/Arith/Int.lean @@ -27,7 +27,7 @@ class HasIntPred {a: Sort u} (x: a) where    prop : concl  /- Proposition with implications: if we find P we can introduce Q in the context -/ -class PropHasImp (x : Prop) where +class PropHasImp (x : Sort u) where    concl : Prop    prop : x → concl @@ -199,7 +199,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 @@ -242,7 +242,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 ()) @@ -259,10 +259,10 @@ 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 +  -- Call omega    Tactic.allGoals do      try do Tactic.Omega.omegaTactic {}      catch _ => diff --git a/backends/lean/Base/Arith/Scalar.lean b/backends/lean/Base/Arith/Scalar.lean index 86b2e216..ecc5acaf 100644 --- a/backends/lean/Base/Arith/Scalar.lean +++ b/backends/lean/Base/Arith/Scalar.lean @@ -8,30 +8,31 @@ 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 {} +               -- Simprocs +               #[] +               -- 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 +82,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  | 
