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/Int.lean | |
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/Int.lean | 10 |
1 files changed, 5 insertions, 5 deletions
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 _ => |