From 7763a8ef8d5190fad39e9e677c5f44c536973655 Mon Sep 17 00:00:00 2001 From: Son Ho Date: Wed, 12 Jun 2024 14:45:58 +0200 Subject: Add the Simp.Config to the simp wrappers --- backends/lean/Base/Arith/Int.lean | 6 ++-- backends/lean/Base/Arith/Scalar.lean | 47 +++++++++++++++---------------- backends/lean/Base/Progress/Progress.lean | 4 +-- backends/lean/Base/Utils.lean | 16 +++++------ 4 files changed, 36 insertions(+), 37 deletions(-) (limited to 'backends') 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..c2e4e24e 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 diff --git a/backends/lean/Base/Progress/Progress.lean b/backends/lean/Base/Progress/Progress.lean index f2a56e50..39a48044 100644 --- a/backends/lean/Base/Progress/Progress.lean +++ b/backends/lean/Base/Progress/Progress.lean @@ -135,7 +135,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 +144,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..6ee854cc 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,28 +731,28 @@ 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 -- cgit v1.2.3 From 191927e2fd21cf94501ac4f4968b09d110043d33 Mon Sep 17 00:00:00 2001 From: Son Ho Date: Wed, 12 Jun 2024 14:46:21 +0200 Subject: Add wrappers around the rewriter --- backends/lean/Base/Utils.lean | 61 +++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 61 insertions(+) (limited to 'backends') diff --git a/backends/lean/Base/Utils.lean b/backends/lean/Base/Utils.lean index 6ee854cc..4be46400 100644 --- a/backends/lean/Base/Utils.lean +++ b/backends/lean/Base/Utils.lean @@ -756,4 +756,65 @@ def simpAll (config : Simp.Config) (declsToUnfold : List Name) (thms : List Name -- 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 -- cgit v1.2.3 From f5deac2b0f42e2a87fc26da50c902729e0ed1039 Mon Sep 17 00:00:00 2001 From: Son Ho Date: Wed, 12 Jun 2024 14:46:38 +0200 Subject: Slightly simplify the progress tactic --- backends/lean/Base/Progress/Progress.lean | 12 ++++-------- 1 file changed, 4 insertions(+), 8 deletions(-) (limited to 'backends') diff --git a/backends/lean/Base/Progress/Progress.lean b/backends/lean/Base/Progress/Progress.lean index 39a48044..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 -- cgit v1.2.3 From 1018aa1af83e639a6b41b5650bf3b717e7f8de68 Mon Sep 17 00:00:00 2001 From: Son Ho Date: Wed, 12 Jun 2024 14:46:52 +0200 Subject: Deactivate the coercion from Nat to Scalar --- backends/lean/Base/Arith/Scalar.lean | 10 ++++++++++ backends/lean/Base/Primitives/Scalar.lean | 7 +++++++ 2 files changed, 17 insertions(+) (limited to 'backends') diff --git a/backends/lean/Base/Arith/Scalar.lean b/backends/lean/Base/Arith/Scalar.lean index c2e4e24e..8793713b 100644 --- a/backends/lean/Base/Arith/Scalar.lean +++ b/backends/lean/Base/Arith/Scalar.lean @@ -80,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) : -- cgit v1.2.3 From c8272aeea205ca9cb36e22757473ca2a931a4933 Mon Sep 17 00:00:00 2001 From: Son Ho Date: Wed, 12 Jun 2024 14:52:34 +0200 Subject: Update the scalar notation for the Lean backend --- backends/lean/Base/Primitives/Scalar.lean | 50 +++++++++++++++---------------- 1 file changed, 25 insertions(+), 25 deletions(-) (limited to 'backends') diff --git a/backends/lean/Base/Primitives/Scalar.lean b/backends/lean/Base/Primitives/Scalar.lean index 157ade2c..5f14a14f 100644 --- a/backends/lean/Base/Primitives/Scalar.lean +++ b/backends/lean/Base/Primitives/Scalar.lean @@ -313,13 +313,13 @@ theorem Scalar.bound_suffices (ty : ScalarTy) (x : Int) : apply And.intro <;> have hmin := Scalar.cMin_bound ty <;> have hmax := Scalar.cMax_bound ty <;> linarith /- [match_pattern] attribute: allows to us `Scalar.ofIntCore` inside of patterns. - This is particularly useful once we introduce notations like `#u32` (which + This is particularly useful once we introduce notations like `u32` (which desugards to `Scalar.ofIntCore`) as it allows to write expressions like this: Example: ``` match x with - | 0#u32 => ... - | 1#u32 => ... + | 0u32 => ... + | 1u32 => ... | ... ``` -/ @@ -328,7 +328,7 @@ theorem Scalar.bound_suffices (ty : ScalarTy) (x : Int) : { val := x, hmin := h.left, hmax := h.right } -- The definitions below are used later to introduce nice syntax for constants, --- like `1#u32`. We are reusing the technique described here: https://leanprover.zulipchat.com/#narrow/stream/270676-lean4/topic/Different.20elaboration.20inside.2Foutside.20of.20match.20patterns/near/425455284 +-- like `1u32`. We are reusing the technique described here: https://leanprover.zulipchat.com/#narrow/stream/270676-lean4/topic/Different.20elaboration.20inside.2Foutside.20of.20match.20patterns/near/425455284 class InBounds (ty : ScalarTy) (x : Int) := hInBounds : Scalar.cMin ty ≤ x ∧ x ≤ Scalar.cMax ty @@ -1281,38 +1281,38 @@ def U128.ofIntCore := @Scalar.ofIntCore .U128 @[match_pattern] abbrev U64.ofInt := @Scalar.ofInt .U64 @[match_pattern] abbrev U128.ofInt := @Scalar.ofInt .U128 -postfix:max "#isize" => Isize.ofInt -postfix:max "#i8" => I8.ofInt -postfix:max "#i16" => I16.ofInt -postfix:max "#i32" => I32.ofInt -postfix:max "#i64" => I64.ofInt -postfix:max "#i128" => I128.ofInt -postfix:max "#usize" => Usize.ofInt -postfix:max "#u8" => U8.ofInt -postfix:max "#u16" => U16.ofInt -postfix:max "#u32" => U32.ofInt -postfix:max "#u64" => U64.ofInt -postfix:max "#u128" => U128.ofInt +postfix:max "isize" => Isize.ofInt +postfix:max "i8" => I8.ofInt +postfix:max "i16" => I16.ofInt +postfix:max "i32" => I32.ofInt +postfix:max "i64" => I64.ofInt +postfix:max "i128" => I128.ofInt +postfix:max "usize" => Usize.ofInt +postfix:max "u8" => U8.ofInt +postfix:max "u16" => U16.ofInt +postfix:max "u32" => U32.ofInt +postfix:max "u64" => U64.ofInt +postfix:max "u128" => U128.ofInt /- Testing the notations -/ -example := 0#u32 -example := 1#u32 -example := 1#i32 -example := 0#isize -example := (-1)#isize +example := 0u32 +example := 1u32 +example := 1i32 +example := 0isize +example := (-1)isize example (x : U32) : Bool := match x with - | 0#u32 => true + | 0u32 => true | _ => false example (x : U32) : Bool := match x with - | 1#u32 => true + | 1u32 => true | _ => false example (x : I32) : Bool := match x with - | (-1)#i32 => true + | (-1)i32 => true | _ => false -- Notation for pattern matching @@ -1334,7 +1334,7 @@ example {ty} (x : Scalar ty) : Bool := | _ => false -- Testing the notations -example : Result Usize := 0#usize + 1#usize +example : Result Usize := 0usize + 1usize -- TODO: factor those lemmas out @[simp] theorem Scalar.ofInt_val_eq {ty} (h : Scalar.min ty ≤ x ∧ x ≤ Scalar.max ty) : (Scalar.ofIntCore x h).val = x := by -- cgit v1.2.3 From fe59d3afe131e9f83cbdf73b1c8089ad090d92fb Mon Sep 17 00:00:00 2001 From: Son Ho Date: Wed, 12 Jun 2024 16:18:05 +0200 Subject: Update CoreConvertNum.lean --- backends/lean/Base/Primitives/CoreConvertNum.lean | 44 +++++++++++------------ 1 file changed, 22 insertions(+), 22 deletions(-) (limited to 'backends') diff --git a/backends/lean/Base/Primitives/CoreConvertNum.lean b/backends/lean/Base/Primitives/CoreConvertNum.lean index eb456a96..bc3db8b2 100644 --- a/backends/lean/Base/Primitives/CoreConvertNum.lean +++ b/backends/lean/Base/Primitives/CoreConvertNum.lean @@ -18,40 +18,40 @@ namespace num -- core.convert.num -- Conversions def FromUsizeBool.from (b : Bool) : Usize := - if b then 1#usize else 0#usize + if b then 1usize else 0usize def FromU8Bool.from (b : Bool) : U8 := - if b then 1#u8 else 0#u8 + if b then 1u8 else 0u8 def FromU16Bool.from (b : Bool) : U16 := - if b then 1#u16 else 0#u16 + if b then 1u16 else 0u16 def FromU32Bool.from (b : Bool) : U32 := - if b then 1#u32 else 0#u32 + if b then 1u32 else 0u32 def FromU64Bool.from (b : Bool) : U64 := - if b then 1#u64 else 0#u64 + if b then 1u64 else 0u64 def FromU128Bool.from (b : Bool) : U128 := - if b then 1#u128 else 0#u128 + if b then 1u128 else 0u128 def FromIsizeBool.from (b : Bool) : Isize := - if b then 1#isize else 0#isize + if b then 1isize else 0isize def FromI8Bool.from (b : Bool) : I8 := - if b then 1#i8 else 0#i8 + if b then 1i8 else 0i8 def FromI16Bool.from (b : Bool) : I16 := - if b then 1#i16 else 0#i16 + if b then 1i16 else 0i16 def FromI32Bool.from (b : Bool) : I32 := - if b then 1#i32 else 0#i32 + if b then 1i32 else 0i32 def FromI64Bool.from (b : Bool) : I64 := - if b then 1#i64 else 0#i64 + if b then 1i64 else 0i64 def FromI128Bool.from (b : Bool) : I128 := - if b then 1#i128 else 0#i128 + if b then 1i128 else 0i128 def FromUsizeU8.from (x : U8) : Usize := ⟨ x.val, by scalar_tac, by scalar_tac ⟩ @@ -259,18 +259,18 @@ def FromI128I128 : core.convert.From I128 I128 := { } -- to_le_bytes -def core.num.U8.to_le_bytes (x : U8) : Array U8 1#usize := sorry -def core.num.U16.to_le_bytes (x : U16) : Array U8 2#usize := sorry -def core.num.U32.to_le_bytes (x : U32) : Array U8 4#usize := sorry -def core.num.U64.to_le_bytes (x : U64) : Array U8 8#usize := sorry -def core.num.U128.to_le_bytes (x : U128) : Array U8 128#usize := sorry +def core.num.U8.to_le_bytes (x : U8) : Array U8 1usize := sorry +def core.num.U16.to_le_bytes (x : U16) : Array U8 2usize := sorry +def core.num.U32.to_le_bytes (x : U32) : Array U8 4usize := sorry +def core.num.U64.to_le_bytes (x : U64) : Array U8 8usize := sorry +def core.num.U128.to_le_bytes (x : U128) : Array U8 128usize := sorry -- to_be_bytes -def core.num.U8.to_be_bytes (x : U8) : Array U8 1#usize := sorry -def core.num.U16.to_be_bytes (x : U16) : Array U8 2#usize := sorry -def core.num.U32.to_be_bytes (x : U32) : Array U8 4#usize := sorry -def core.num.U64.to_be_bytes (x : U64) : Array U8 8#usize := sorry -def core.num.U128.to_be_bytes (x : U128) : Array U8 128#usize := sorry +def core.num.U8.to_be_bytes (x : U8) : Array U8 1usize := sorry +def core.num.U16.to_be_bytes (x : U16) : Array U8 2usize := sorry +def core.num.U32.to_be_bytes (x : U32) : Array U8 4usize := sorry +def core.num.U64.to_be_bytes (x : U64) : Array U8 8usize := sorry +def core.num.U128.to_be_bytes (x : U128) : Array U8 128usize := sorry end core.convert -- cgit v1.2.3 From dd2b973a86680308807d7f26aff81d3310550f84 Mon Sep 17 00:00:00 2001 From: Son Ho Date: Wed, 12 Jun 2024 18:20:23 +0200 Subject: Revert "Update CoreConvertNum.lean" This reverts commit fe59d3afe131e9f83cbdf73b1c8089ad090d92fb. --- backends/lean/Base/Primitives/CoreConvertNum.lean | 44 +++++++++++------------ 1 file changed, 22 insertions(+), 22 deletions(-) (limited to 'backends') diff --git a/backends/lean/Base/Primitives/CoreConvertNum.lean b/backends/lean/Base/Primitives/CoreConvertNum.lean index bc3db8b2..eb456a96 100644 --- a/backends/lean/Base/Primitives/CoreConvertNum.lean +++ b/backends/lean/Base/Primitives/CoreConvertNum.lean @@ -18,40 +18,40 @@ namespace num -- core.convert.num -- Conversions def FromUsizeBool.from (b : Bool) : Usize := - if b then 1usize else 0usize + if b then 1#usize else 0#usize def FromU8Bool.from (b : Bool) : U8 := - if b then 1u8 else 0u8 + if b then 1#u8 else 0#u8 def FromU16Bool.from (b : Bool) : U16 := - if b then 1u16 else 0u16 + if b then 1#u16 else 0#u16 def FromU32Bool.from (b : Bool) : U32 := - if b then 1u32 else 0u32 + if b then 1#u32 else 0#u32 def FromU64Bool.from (b : Bool) : U64 := - if b then 1u64 else 0u64 + if b then 1#u64 else 0#u64 def FromU128Bool.from (b : Bool) : U128 := - if b then 1u128 else 0u128 + if b then 1#u128 else 0#u128 def FromIsizeBool.from (b : Bool) : Isize := - if b then 1isize else 0isize + if b then 1#isize else 0#isize def FromI8Bool.from (b : Bool) : I8 := - if b then 1i8 else 0i8 + if b then 1#i8 else 0#i8 def FromI16Bool.from (b : Bool) : I16 := - if b then 1i16 else 0i16 + if b then 1#i16 else 0#i16 def FromI32Bool.from (b : Bool) : I32 := - if b then 1i32 else 0i32 + if b then 1#i32 else 0#i32 def FromI64Bool.from (b : Bool) : I64 := - if b then 1i64 else 0i64 + if b then 1#i64 else 0#i64 def FromI128Bool.from (b : Bool) : I128 := - if b then 1i128 else 0i128 + if b then 1#i128 else 0#i128 def FromUsizeU8.from (x : U8) : Usize := ⟨ x.val, by scalar_tac, by scalar_tac ⟩ @@ -259,18 +259,18 @@ def FromI128I128 : core.convert.From I128 I128 := { } -- to_le_bytes -def core.num.U8.to_le_bytes (x : U8) : Array U8 1usize := sorry -def core.num.U16.to_le_bytes (x : U16) : Array U8 2usize := sorry -def core.num.U32.to_le_bytes (x : U32) : Array U8 4usize := sorry -def core.num.U64.to_le_bytes (x : U64) : Array U8 8usize := sorry -def core.num.U128.to_le_bytes (x : U128) : Array U8 128usize := sorry +def core.num.U8.to_le_bytes (x : U8) : Array U8 1#usize := sorry +def core.num.U16.to_le_bytes (x : U16) : Array U8 2#usize := sorry +def core.num.U32.to_le_bytes (x : U32) : Array U8 4#usize := sorry +def core.num.U64.to_le_bytes (x : U64) : Array U8 8#usize := sorry +def core.num.U128.to_le_bytes (x : U128) : Array U8 128#usize := sorry -- to_be_bytes -def core.num.U8.to_be_bytes (x : U8) : Array U8 1usize := sorry -def core.num.U16.to_be_bytes (x : U16) : Array U8 2usize := sorry -def core.num.U32.to_be_bytes (x : U32) : Array U8 4usize := sorry -def core.num.U64.to_be_bytes (x : U64) : Array U8 8usize := sorry -def core.num.U128.to_be_bytes (x : U128) : Array U8 128usize := sorry +def core.num.U8.to_be_bytes (x : U8) : Array U8 1#usize := sorry +def core.num.U16.to_be_bytes (x : U16) : Array U8 2#usize := sorry +def core.num.U32.to_be_bytes (x : U32) : Array U8 4#usize := sorry +def core.num.U64.to_be_bytes (x : U64) : Array U8 8#usize := sorry +def core.num.U128.to_be_bytes (x : U128) : Array U8 128#usize := sorry end core.convert -- cgit v1.2.3 From d36736fa4e7eb9f42f35303b8080d17ddbee92d2 Mon Sep 17 00:00:00 2001 From: Son Ho Date: Wed, 12 Jun 2024 18:20:52 +0200 Subject: Revert "Update the scalar notation for the Lean backend" This reverts commit c8272aeea205ca9cb36e22757473ca2a931a4933. --- backends/lean/Base/Primitives/Scalar.lean | 50 +++++++++++++++---------------- 1 file changed, 25 insertions(+), 25 deletions(-) (limited to 'backends') diff --git a/backends/lean/Base/Primitives/Scalar.lean b/backends/lean/Base/Primitives/Scalar.lean index 5f14a14f..157ade2c 100644 --- a/backends/lean/Base/Primitives/Scalar.lean +++ b/backends/lean/Base/Primitives/Scalar.lean @@ -313,13 +313,13 @@ theorem Scalar.bound_suffices (ty : ScalarTy) (x : Int) : apply And.intro <;> have hmin := Scalar.cMin_bound ty <;> have hmax := Scalar.cMax_bound ty <;> linarith /- [match_pattern] attribute: allows to us `Scalar.ofIntCore` inside of patterns. - This is particularly useful once we introduce notations like `u32` (which + This is particularly useful once we introduce notations like `#u32` (which desugards to `Scalar.ofIntCore`) as it allows to write expressions like this: Example: ``` match x with - | 0u32 => ... - | 1u32 => ... + | 0#u32 => ... + | 1#u32 => ... | ... ``` -/ @@ -328,7 +328,7 @@ theorem Scalar.bound_suffices (ty : ScalarTy) (x : Int) : { val := x, hmin := h.left, hmax := h.right } -- The definitions below are used later to introduce nice syntax for constants, --- like `1u32`. We are reusing the technique described here: https://leanprover.zulipchat.com/#narrow/stream/270676-lean4/topic/Different.20elaboration.20inside.2Foutside.20of.20match.20patterns/near/425455284 +-- like `1#u32`. We are reusing the technique described here: https://leanprover.zulipchat.com/#narrow/stream/270676-lean4/topic/Different.20elaboration.20inside.2Foutside.20of.20match.20patterns/near/425455284 class InBounds (ty : ScalarTy) (x : Int) := hInBounds : Scalar.cMin ty ≤ x ∧ x ≤ Scalar.cMax ty @@ -1281,38 +1281,38 @@ def U128.ofIntCore := @Scalar.ofIntCore .U128 @[match_pattern] abbrev U64.ofInt := @Scalar.ofInt .U64 @[match_pattern] abbrev U128.ofInt := @Scalar.ofInt .U128 -postfix:max "isize" => Isize.ofInt -postfix:max "i8" => I8.ofInt -postfix:max "i16" => I16.ofInt -postfix:max "i32" => I32.ofInt -postfix:max "i64" => I64.ofInt -postfix:max "i128" => I128.ofInt -postfix:max "usize" => Usize.ofInt -postfix:max "u8" => U8.ofInt -postfix:max "u16" => U16.ofInt -postfix:max "u32" => U32.ofInt -postfix:max "u64" => U64.ofInt -postfix:max "u128" => U128.ofInt +postfix:max "#isize" => Isize.ofInt +postfix:max "#i8" => I8.ofInt +postfix:max "#i16" => I16.ofInt +postfix:max "#i32" => I32.ofInt +postfix:max "#i64" => I64.ofInt +postfix:max "#i128" => I128.ofInt +postfix:max "#usize" => Usize.ofInt +postfix:max "#u8" => U8.ofInt +postfix:max "#u16" => U16.ofInt +postfix:max "#u32" => U32.ofInt +postfix:max "#u64" => U64.ofInt +postfix:max "#u128" => U128.ofInt /- Testing the notations -/ -example := 0u32 -example := 1u32 -example := 1i32 -example := 0isize -example := (-1)isize +example := 0#u32 +example := 1#u32 +example := 1#i32 +example := 0#isize +example := (-1)#isize example (x : U32) : Bool := match x with - | 0u32 => true + | 0#u32 => true | _ => false example (x : U32) : Bool := match x with - | 1u32 => true + | 1#u32 => true | _ => false example (x : I32) : Bool := match x with - | (-1)i32 => true + | (-1)#i32 => true | _ => false -- Notation for pattern matching @@ -1334,7 +1334,7 @@ example {ty} (x : Scalar ty) : Bool := | _ => false -- Testing the notations -example : Result Usize := 0usize + 1usize +example : Result Usize := 0#usize + 1#usize -- TODO: factor those lemmas out @[simp] theorem Scalar.ofInt_val_eq {ty} (h : Scalar.min ty ≤ x ∧ x ≤ Scalar.max ty) : (Scalar.ofIntCore x h).val = x := by -- cgit v1.2.3