From e2fef1a5c986aff4c9975b1376bcc0fc0bb87940 Mon Sep 17 00:00:00 2001 From: Son Ho Date: Fri, 9 Jun 2023 10:06:43 +0200 Subject: Reorganize a bit the Lean library --- backends/lean/Base/Primitives.lean | 650 +++++++++++++++++++++++++++++++++++++ 1 file changed, 650 insertions(+) create mode 100644 backends/lean/Base/Primitives.lean (limited to 'backends/lean/Base/Primitives.lean') diff --git a/backends/lean/Base/Primitives.lean b/backends/lean/Base/Primitives.lean new file mode 100644 index 00000000..d3de1d10 --- /dev/null +++ b/backends/lean/Base/Primitives.lean @@ -0,0 +1,650 @@ +import Lean +import Lean.Meta.Tactic.Simp +import Init.Data.List.Basic +import Mathlib.Tactic.RunCmd +import Mathlib.Tactic.Linarith + +-------------------- +-- ASSERT COMMAND --Std. +-------------------- + +open Lean Elab Command Term Meta + +syntax (name := assert) "#assert" term: command + +@[command_elab assert] +unsafe +def assertImpl : CommandElab := fun (_stx: Syntax) => do + runTermElabM (fun _ => do + let r ← evalTerm Bool (mkConst ``Bool) _stx[1] + if not r then + logInfo "Assertion failed for: " + logInfo _stx[1] + logError "Expression reduced to false" + pure ()) + +#eval 2 == 2 +#assert (2 == 2) + +------------- +-- PRELUDE -- +------------- + +-- Results & monadic combinators + +inductive Error where + | assertionFailure: Error + | integerOverflow: Error + | divisionByZero: Error + | arrayOutOfBounds: Error + | maximumSizeExceeded: Error + | panic: Error +deriving Repr, BEq + +open Error + +inductive Result (α : Type u) where + | ret (v: α): Result α + | fail (e: Error): Result α +deriving Repr, BEq + +open Result + +instance Result_Inhabited (α : Type u) : Inhabited (Result α) := + Inhabited.mk (fail panic) + +/- HELPERS -/ + +def ret? {α: Type} (r: Result α): Bool := + match r with + | Result.ret _ => true + | Result.fail _ => false + +def massert (b:Bool) : Result Unit := + if b then .ret () else fail assertionFailure + +def eval_global {α: Type} (x: Result α) (_: ret? x): α := + match x with + | Result.fail _ => by contradiction + | Result.ret x => x + +/- DO-DSL SUPPORT -/ + +def bind (x: Result α) (f: α -> Result β) : Result β := + match x with + | ret v => f v + | fail v => fail v + +-- Allows using Result in do-blocks +instance : Bind Result where + bind := bind + +-- Allows using return x in do-blocks +instance : Pure Result where + pure := fun x => ret x + +/- CUSTOM-DSL SUPPORT -/ + +-- Let-binding the Result of a monadic operation is oftentimes not sufficient, +-- because we may need a hypothesis for equational reasoning in the scope. We +-- rely on subtype, and a custom let-binding operator, in effect recreating our +-- own variant of the do-dsl + +def Result.attach {α: Type} (o : Result α): Result { x : α // o = ret x } := + match o with + | .ret x => .ret ⟨x, rfl⟩ + | .fail e => .fail e + +macro "let" e:term " ⟵ " f:term : doElem => + `(doElem| let ⟨$e, h⟩ ← Result.attach $f) + +-- TODO: any way to factorize both definitions? +macro "let" e:term " <-- " f:term : doElem => + `(doElem| let ⟨$e, h⟩ ← Result.attach $f) + +-- We call the hypothesis `h`, in effect making it unavailable to the user +-- (because too much shadowing). But in practice, once can use the French single +-- quote notation (input with f< and f>), where `‹ h ›` finds a suitable +-- hypothesis in the context, this is equivalent to `have x: h := by assumption in x` +#eval do + let y <-- .ret (0: Nat) + let _: y = 0 := by cases ‹ ret 0 = ret y › ; decide + let r: { x: Nat // x = 0 } := ⟨ y, by assumption ⟩ + .ret r + +---------------------- +-- MACHINE INTEGERS -- +---------------------- + +-- We redefine our machine integers types. + +-- For Isize/Usize, we reuse `getNumBits` from `USize`. You cannot reduce `getNumBits` +-- using the simplifier, meaning that proofs do not depend on the compile-time value of +-- USize.size. (Lean assumes 32 or 64-bit platforms, and Rust doesn't really support, at +-- least officially, 16-bit microcontrollers, so this seems like a fine design decision +-- for now.) + +-- Note from Chris Bailey: "If there's more than one salient property of your +-- definition then the subtyping strategy might get messy, and the property part +-- of a subtype is less discoverable by the simplifier or tactics like +-- library_search." So, we will not add refinements on the return values of the +-- operations defined on Primitives, but will rather rely on custom lemmas to +-- invert on possible return values of the primitive operations. + +-- Machine integer constants, done via `ofNatCore`, which requires a proof that +-- the `Nat` fits within the desired integer type. We provide a custom tactic. + +open System.Platform.getNumBits + +-- TODO: is there a way of only importing System.Platform.getNumBits? +-- +@[simp] def size_num_bits : Nat := (System.Platform.getNumBits ()).val + +-- Remark: Lean seems to use < for the comparisons with the upper bounds by convention. +-- We keep the F* convention for now. +@[simp] def Isize.min : Int := - (HPow.hPow 2 (size_num_bits - 1)) +@[simp] def Isize.max : Int := (HPow.hPow 2 (size_num_bits - 1)) - 1 +@[simp] def I8.min : Int := - (HPow.hPow 2 7) +@[simp] def I8.max : Int := HPow.hPow 2 7 - 1 +@[simp] def I16.min : Int := - (HPow.hPow 2 15) +@[simp] def I16.max : Int := HPow.hPow 2 15 - 1 +@[simp] def I32.min : Int := -(HPow.hPow 2 31) +@[simp] def I32.max : Int := HPow.hPow 2 31 - 1 +@[simp] def I64.min : Int := -(HPow.hPow 2 63) +@[simp] def I64.max : Int := HPow.hPow 2 63 - 1 +@[simp] def I128.min : Int := -(HPow.hPow 2 127) +@[simp] def I128.max : Int := HPow.hPow 2 127 - 1 +@[simp] def Usize.min : Int := 0 +@[simp] def Usize.max : Int := HPow.hPow 2 size_num_bits - 1 +@[simp] def U8.min : Int := 0 +@[simp] def U8.max : Int := HPow.hPow 2 8 - 1 +@[simp] def U16.min : Int := 0 +@[simp] def U16.max : Int := HPow.hPow 2 16 - 1 +@[simp] def U32.min : Int := 0 +@[simp] def U32.max : Int := HPow.hPow 2 32 - 1 +@[simp] def U64.min : Int := 0 +@[simp] def U64.max : Int := HPow.hPow 2 64 - 1 +@[simp] def U128.min : Int := 0 +@[simp] def U128.max : Int := HPow.hPow 2 128 - 1 + +#assert (I8.min == -128) +#assert (I8.max == 127) +#assert (I16.min == -32768) +#assert (I16.max == 32767) +#assert (I32.min == -2147483648) +#assert (I32.max == 2147483647) +#assert (I64.min == -9223372036854775808) +#assert (I64.max == 9223372036854775807) +#assert (I128.min == -170141183460469231731687303715884105728) +#assert (I128.max == 170141183460469231731687303715884105727) +#assert (U8.min == 0) +#assert (U8.max == 255) +#assert (U16.min == 0) +#assert (U16.max == 65535) +#assert (U32.min == 0) +#assert (U32.max == 4294967295) +#assert (U64.min == 0) +#assert (U64.max == 18446744073709551615) +#assert (U128.min == 0) +#assert (U128.max == 340282366920938463463374607431768211455) + +inductive ScalarTy := +| Isize +| I8 +| I16 +| I32 +| I64 +| I128 +| Usize +| U8 +| U16 +| U32 +| U64 +| U128 + +def Scalar.min (ty : ScalarTy) : Int := + match ty with + | .Isize => Isize.min + | .I8 => I8.min + | .I16 => I16.min + | .I32 => I32.min + | .I64 => I64.min + | .I128 => I128.min + | .Usize => Usize.min + | .U8 => U8.min + | .U16 => U16.min + | .U32 => U32.min + | .U64 => U64.min + | .U128 => U128.min + +def Scalar.max (ty : ScalarTy) : Int := + match ty with + | .Isize => Isize.max + | .I8 => I8.max + | .I16 => I16.max + | .I32 => I32.max + | .I64 => I64.max + | .I128 => I128.max + | .Usize => Usize.max + | .U8 => U8.max + | .U16 => U16.max + | .U32 => U32.max + | .U64 => U64.max + | .U128 => U128.max + +-- "Conservative" bounds +-- We use those because we can't compare to the isize bounds (which can't +-- reduce at compile-time). Whenever we perform an arithmetic operation like +-- addition we need to check that the result is in bounds: we first compare +-- to the conservative bounds, which reduce, then compare to the real bounds. +-- This is useful for the various #asserts that we want to reduce at +-- type-checking time. +def Scalar.cMin (ty : ScalarTy) : Int := + match ty with + | .Isize => I32.min + | _ => Scalar.min ty + +def Scalar.cMax (ty : ScalarTy) : Int := + match ty with + | .Isize => I32.max + | .Usize => U32.max + | _ => Scalar.max ty + +theorem Scalar.cMin_bound ty : Scalar.min ty ≤ Scalar.cMin ty := by + cases ty <;> simp [Scalar.min, Scalar.max, Scalar.cMin, Scalar.cMax] at * + cases System.Platform.numBits_eq <;> + unfold System.Platform.numBits at * <;> + simp [*] + +theorem Scalar.cMax_bound ty : Scalar.cMax ty ≤ Scalar.max ty := by + cases ty <;> simp [Scalar.min, Scalar.max, Scalar.cMin, Scalar.cMax] at * <;> + cases System.Platform.numBits_eq <;> + unfold System.Platform.numBits at * <;> + simp [*] + +theorem Scalar.cMin_suffices ty (h : Scalar.cMin ty ≤ x) : Scalar.min ty ≤ x := by + cases ty <;> simp [Scalar.min, Scalar.max, Scalar.cMin, Scalar.cMax] at * <;> + cases System.Platform.numBits_eq <;> + unfold System.Platform.numBits at * <;> + simp [*] at * + -- TODO: I would have expected terms like `-(1 + 1) ^ 63` to be simplified + linarith + +theorem Scalar.cMax_suffices ty (h : x ≤ Scalar.cMax ty) : x ≤ Scalar.max ty := by + cases ty <;> simp [Scalar.min, Scalar.max, Scalar.cMin, Scalar.cMax] at * <;> + cases System.Platform.numBits_eq <;> + unfold System.Platform.numBits at * <;> + simp [*] at * <;> + -- TODO: I would have expected terms like `-(1 + 1) ^ 63` to be simplified + linarith + +structure Scalar (ty : ScalarTy) where + val : Int + hmin : Scalar.min ty ≤ val + hmax : val ≤ Scalar.max ty +deriving Repr + +theorem Scalar.bound_suffices (ty : ScalarTy) (x : Int) : + Scalar.cMin ty ≤ x ∧ x ≤ Scalar.cMax ty -> + Scalar.min ty ≤ x ∧ x ≤ Scalar.max ty + := + λ h => by + apply And.intro <;> have hmin := Scalar.cMin_bound ty <;> have hmax := Scalar.cMax_bound ty <;> linarith + +def Scalar.ofIntCore {ty : ScalarTy} (x : Int) + (hmin : Scalar.min ty ≤ x) (hmax : x ≤ Scalar.max ty) : Scalar ty := + { val := x, hmin := hmin, hmax := hmax } + +def Scalar.ofInt {ty : ScalarTy} (x : Int) + (h : Scalar.min ty ≤ x ∧ x ≤ Scalar.max ty) : Scalar ty := + -- Remark: we initially wrote: + -- let ⟨ hmin, hmax ⟩ := h + -- Scalar.ofIntCore x hmin hmax + -- We updated to the line below because a similar pattern in `Scalar.tryMk` + -- made reduction block. Both versions seem to work for `Scalar.ofInt`, though. + -- TODO: investigate + Scalar.ofIntCore x h.left h.right + +@[simp] def Scalar.check_bounds (ty : ScalarTy) (x : Int) : Bool := + (Scalar.cMin ty ≤ x || Scalar.min ty ≤ x) ∧ (x ≤ Scalar.cMax ty || x ≤ Scalar.max ty) + +theorem Scalar.check_bounds_prop {ty : ScalarTy} {x : Int} (h: Scalar.check_bounds ty x) : + Scalar.min ty ≤ x ∧ x ≤ Scalar.max ty := by + simp at * + have ⟨ hmin, hmax ⟩ := h + have hbmin := Scalar.cMin_bound ty + have hbmax := Scalar.cMax_bound ty + cases hmin <;> cases hmax <;> apply And.intro <;> linarith + +-- Further thoughts: look at what has been done here: +-- https://github.com/leanprover-community/mathlib4/blob/master/Mathlib/Data/Fin/Basic.lean +-- and +-- https://github.com/leanprover-community/mathlib4/blob/master/Mathlib/Data/UInt.lean +-- which both contain a fair amount of reasoning already! +def Scalar.tryMk (ty : ScalarTy) (x : Int) : Result (Scalar ty) := + if h:Scalar.check_bounds ty x then + -- If we do: + -- ``` + -- let ⟨ hmin, hmax ⟩ := (Scalar.check_bounds_prop h) + -- Scalar.ofIntCore x hmin hmax + -- ``` + -- then normalization blocks (for instance, some proofs which use reflexivity fail). + -- However, the version below doesn't block reduction (TODO: investigate): + return Scalar.ofInt x (Scalar.check_bounds_prop h) + else fail integerOverflow + +def Scalar.neg {ty : ScalarTy} (x : Scalar ty) : Result (Scalar ty) := Scalar.tryMk ty (- x.val) + +def Scalar.div {ty : ScalarTy} (x : Scalar ty) (y : Scalar ty) : Result (Scalar ty) := + if y.val != 0 then Scalar.tryMk ty (x.val / y.val) else fail divisionByZero + +-- Our custom remainder operation, which satisfies the semantics of Rust +-- TODO: is there a better way? +def scalar_rem (x y : Int) : Int := + if 0 ≤ x then |x| % |y| + else - (|x| % |y|) + +-- Our custom division operation, which satisfies the semantics of Rust +-- TODO: is there a better way? +def scalar_div (x y : Int) : Int := + if 0 ≤ x && 0 ≤ y then |x| / |y| + else if 0 ≤ x && y < 0 then - (|x| / |y|) + else if x < 0 && 0 ≤ y then - (|x| / |y|) + else |x| / |y| + +-- Checking that the remainder operation is correct +#assert scalar_rem 1 2 = 1 +#assert scalar_rem (-1) 2 = -1 +#assert scalar_rem 1 (-2) = 1 +#assert scalar_rem (-1) (-2) = -1 +#assert scalar_rem 7 3 = (1:Int) +#assert scalar_rem (-7) 3 = -1 +#assert scalar_rem 7 (-3) = 1 +#assert scalar_rem (-7) (-3) = -1 + +-- Checking that the division operation is correct +#assert scalar_div 3 2 = 1 +#assert scalar_div (-3) 2 = -1 +#assert scalar_div 3 (-2) = -1 +#assert scalar_div (-3) (-2) = 1 +#assert scalar_div 7 3 = 2 +#assert scalar_div (-7) 3 = -2 +#assert scalar_div 7 (-3) = -2 +#assert scalar_div (-7) (-3) = 2 + +def Scalar.rem {ty : ScalarTy} (x : Scalar ty) (y : Scalar ty) : Result (Scalar ty) := + if y.val != 0 then Scalar.tryMk ty (x.val % y.val) else fail divisionByZero + +def Scalar.add {ty : ScalarTy} (x : Scalar ty) (y : Scalar ty) : Result (Scalar ty) := + Scalar.tryMk ty (x.val + y.val) + +def Scalar.sub {ty : ScalarTy} (x : Scalar ty) (y : Scalar ty) : Result (Scalar ty) := + Scalar.tryMk ty (x.val - y.val) + +def Scalar.mul {ty : ScalarTy} (x : Scalar ty) (y : Scalar ty) : Result (Scalar ty) := + Scalar.tryMk ty (x.val * y.val) + +-- TODO: instances of +, -, * etc. for scalars + +-- Cast an integer from a [src_ty] to a [tgt_ty] +-- TODO: check the semantics of casts in Rust +def Scalar.cast {src_ty : ScalarTy} (tgt_ty : ScalarTy) (x : Scalar src_ty) : Result (Scalar tgt_ty) := + Scalar.tryMk tgt_ty x.val + +-- The scalar types +-- We declare the definitions as reducible so that Lean can unfold them (useful +-- for type class resolution for instance). +@[reducible] def Isize := Scalar .Isize +@[reducible] def I8 := Scalar .I8 +@[reducible] def I16 := Scalar .I16 +@[reducible] def I32 := Scalar .I32 +@[reducible] def I64 := Scalar .I64 +@[reducible] def I128 := Scalar .I128 +@[reducible] def Usize := Scalar .Usize +@[reducible] def U8 := Scalar .U8 +@[reducible] def U16 := Scalar .U16 +@[reducible] def U32 := Scalar .U32 +@[reducible] def U64 := Scalar .U64 +@[reducible] def U128 := Scalar .U128 + +-- TODO: below: not sure this is the best way. +-- Should we rather overload operations like +, -, etc.? +-- Also, it is possible to automate the generation of those definitions +-- with macros (but would it be a good idea? It would be less easy to +-- read the file, which is not supposed to change a lot) + +-- Negation + +/-- +Remark: there is no heterogeneous negation in the Lean prelude: we thus introduce +one here. + +The notation typeclass for heterogeneous addition. +This enables the notation `- a : β` where `a : α`. +-/ +class HNeg (α : Type u) (β : outParam (Type v)) where + /-- `- a` computes the negation of `a`. + The meaning of this notation is type-dependent. -/ + hNeg : α → β + +prefix:75 "-" => HNeg.hNeg + +instance : HNeg Isize (Result Isize) where hNeg x := Scalar.neg x +instance : HNeg I8 (Result I8) where hNeg x := Scalar.neg x +instance : HNeg I16 (Result I16) where hNeg x := Scalar.neg x +instance : HNeg I32 (Result I32) where hNeg x := Scalar.neg x +instance : HNeg I64 (Result I64) where hNeg x := Scalar.neg x +instance : HNeg I128 (Result I128) where hNeg x := Scalar.neg x + +-- Addition +instance {ty} : HAdd (Scalar ty) (Scalar ty) (Result (Scalar ty)) where + hAdd x y := Scalar.add x y + +-- Substraction +instance {ty} : HSub (Scalar ty) (Scalar ty) (Result (Scalar ty)) where + hSub x y := Scalar.sub x y + +-- Multiplication +instance {ty} : HMul (Scalar ty) (Scalar ty) (Result (Scalar ty)) where + hMul x y := Scalar.mul x y + +-- Division +instance {ty} : HDiv (Scalar ty) (Scalar ty) (Result (Scalar ty)) where + hDiv x y := Scalar.div x y + +-- Remainder +instance {ty} : HMod (Scalar ty) (Scalar ty) (Result (Scalar ty)) where + hMod x y := Scalar.rem x y + +-- ofIntCore +-- TODO: typeclass? +def Isize.ofIntCore := @Scalar.ofIntCore .Isize +def I8.ofIntCore := @Scalar.ofIntCore .I8 +def I16.ofIntCore := @Scalar.ofIntCore .I16 +def I32.ofIntCore := @Scalar.ofIntCore .I32 +def I64.ofIntCore := @Scalar.ofIntCore .I64 +def I128.ofIntCore := @Scalar.ofIntCore .I128 +def Usize.ofIntCore := @Scalar.ofIntCore .Usize +def U8.ofIntCore := @Scalar.ofIntCore .U8 +def U16.ofIntCore := @Scalar.ofIntCore .U16 +def U32.ofIntCore := @Scalar.ofIntCore .U32 +def U64.ofIntCore := @Scalar.ofIntCore .U64 +def U128.ofIntCore := @Scalar.ofIntCore .U128 + +-- ofInt +-- TODO: typeclass? +def Isize.ofInt := @Scalar.ofInt .Isize +def I8.ofInt := @Scalar.ofInt .I8 +def I16.ofInt := @Scalar.ofInt .I16 +def I32.ofInt := @Scalar.ofInt .I32 +def I64.ofInt := @Scalar.ofInt .I64 +def I128.ofInt := @Scalar.ofInt .I128 +def Usize.ofInt := @Scalar.ofInt .Usize +def U8.ofInt := @Scalar.ofInt .U8 +def U16.ofInt := @Scalar.ofInt .U16 +def U32.ofInt := @Scalar.ofInt .U32 +def U64.ofInt := @Scalar.ofInt .U64 +def U128.ofInt := @Scalar.ofInt .U128 + +-- Comparisons +instance {ty} : LT (Scalar ty) where + lt a b := LT.lt a.val b.val + +instance {ty} : LE (Scalar ty) where le a b := LE.le a.val b.val + +instance Scalar.decLt {ty} (a b : Scalar ty) : Decidable (LT.lt a b) := Int.decLt .. +instance Scalar.decLe {ty} (a b : Scalar ty) : Decidable (LE.le a b) := Int.decLe .. + +theorem Scalar.eq_of_val_eq {ty} : ∀ {i j : Scalar ty}, Eq i.val j.val → Eq i j + | ⟨_, _, _⟩, ⟨_, _, _⟩, rfl => rfl + +theorem Scalar.val_eq_of_eq {ty} {i j : Scalar ty} (h : Eq i j) : Eq i.val j.val := + h ▸ rfl + +theorem Scalar.ne_of_val_ne {ty} {i j : Scalar ty} (h : Not (Eq i.val j.val)) : Not (Eq i j) := + fun h' => absurd (val_eq_of_eq h') h + +instance (ty : ScalarTy) : DecidableEq (Scalar ty) := + fun i j => + match decEq i.val j.val with + | isTrue h => isTrue (Scalar.eq_of_val_eq h) + | isFalse h => isFalse (Scalar.ne_of_val_ne h) + +def Scalar.toInt {ty} (n : Scalar ty) : Int := n.val + +-- Tactic to prove that integers are in bounds +syntax "intlit" : tactic + +macro_rules + | `(tactic| intlit) => `(tactic| apply Scalar.bound_suffices ; decide) + +-- -- We now define a type class that subsumes the various machine integer types, so +-- -- as to write a concise definition for scalar_cast, rather than exhaustively +-- -- enumerating all of the possible pairs. We remark that Rust has sane semantics +-- -- and fails if a cast operation would involve a truncation or modulo. + +-- class MachineInteger (t: Type) where +-- size: Nat +-- val: t -> Fin size +-- ofNatCore: (n:Nat) -> LT.lt n size -> t + +-- set_option hygiene false in +-- run_cmd +-- for typeName in [`UInt8, `UInt16, `UInt32, `UInt64, `USize].map Lean.mkIdent do +-- Lean.Elab.Command.elabCommand (← `( +-- namespace $typeName +-- instance: MachineInteger $typeName where +-- size := size +-- val := val +-- ofNatCore := ofNatCore +-- end $typeName +-- )) + +-- -- Aeneas only instantiates the destination type (`src` is implicit). We rely on +-- -- Lean to infer `src`. + +-- def scalar_cast { src: Type } (dst: Type) [ MachineInteger src ] [ MachineInteger dst ] (x: src): Result dst := +-- if h: MachineInteger.val x < MachineInteger.size dst then +-- .ret (MachineInteger.ofNatCore (MachineInteger.val x).val h) +-- else +-- .fail integerOverflow + +------------- +-- VECTORS -- +------------- + +def Vec (α : Type u) := { l : List α // List.length l ≤ Usize.max } + +def vec_new (α : Type u): Vec α := ⟨ [], by apply Scalar.cMax_suffices .Usize; simp ⟩ + +def vec_len (α : Type u) (v : Vec α) : Usize := + let ⟨ v, l ⟩ := v + Usize.ofIntCore (List.length v) (by simp [Scalar.min]) l + +def vec_push_fwd (α : Type u) (_ : Vec α) (_ : α) : Unit := () + +def vec_push_back (α : Type u) (v : Vec α) (x : α) : Result (Vec α) + := + let nlen := List.length v.val + 1 + if h : nlen ≤ U32.max || nlen ≤ Usize.max then + have h : nlen ≤ Usize.max := by + simp at * + cases System.Platform.numBits_eq <;> + unfold System.Platform.numBits at * <;> + simp [*] at * <;> + try assumption + cases h <;> + linarith + return ⟨ List.concat v.val x, by simp at *; assumption ⟩ + else + fail maximumSizeExceeded + +def vec_insert_fwd (α : Type u) (v: Vec α) (i: Usize) (_: α): Result Unit := + if i.val < List.length v.val then + .ret () + else + .fail arrayOutOfBounds + +def vec_insert_back (α : Type u) (v: Vec α) (i: Usize) (x: α): Result (Vec α) := + if i.val < List.length v.val then + -- TODO: maybe we should redefine a list library which uses integers + -- (instead of natural numbers) + let i := i.val.toNat + .ret ⟨ List.set v.val i x, by + have h: List.length v.val ≤ Usize.max := v.property + simp [*] at * + assumption + ⟩ + else + .fail arrayOutOfBounds + +def vec_index_to_fin {α : Type u} {v: Vec α} {i: Usize} (h : i.val < List.length v.val) : + Fin (List.length v.val) := + let j := i.val.toNat + let h: j < List.length v.val := by + have heq := @Int.toNat_lt (List.length v.val) i.val i.hmin + apply heq.mpr + assumption + ⟨j, h⟩ + +def vec_index_fwd (α : Type u) (v: Vec α) (i: Usize): Result α := + if h: i.val < List.length v.val then + let i := vec_index_to_fin h + .ret (List.get v.val i) + else + .fail arrayOutOfBounds + +def vec_index_back (α : Type u) (v: Vec α) (i: Usize) (_: α): Result Unit := + if i.val < List.length v.val then + .ret () + else + .fail arrayOutOfBounds + +def vec_index_mut_fwd (α : Type u) (v: Vec α) (i: Usize): Result α := + if h: i.val < List.length v.val then + let i := vec_index_to_fin h + .ret (List.get v.val i) + else + .fail arrayOutOfBounds + +def vec_index_mut_back (α : Type u) (v: Vec α) (i: Usize) (x: α): Result (Vec α) := + if h: i.val < List.length v.val then + let i := vec_index_to_fin h + .ret ⟨ List.set v.val i x, by + have h: List.length v.val ≤ Usize.max := v.property + simp [*] at * + assumption + ⟩ + else + .fail arrayOutOfBounds + +---------- +-- MISC -- +---------- + +@[simp] def mem_replace_fwd (a : Type) (x : a) (_ : a) : a := x +@[simp] def mem_replace_back (a : Type) (_ : a) (y : a) : a := y + +/-- Aeneas-translated function -- useful to reduce non-recursive definitions. + Use with `simp [ aeneas ]` -/ +register_simp_attr aeneas -- cgit v1.2.3 From c034a7ea1335705ca1e1a7461fac257df6757d57 Mon Sep 17 00:00:00 2001 From: Son Ho Date: Fri, 9 Jun 2023 16:07:39 +0200 Subject: Start working on extrinsic proofs of termination --- backends/lean/Base/Primitives.lean | 29 ++++++++++++++++++++++------- 1 file changed, 22 insertions(+), 7 deletions(-) (limited to 'backends/lean/Base/Primitives.lean') diff --git a/backends/lean/Base/Primitives.lean b/backends/lean/Base/Primitives.lean index d3de1d10..85e088fc 100644 --- a/backends/lean/Base/Primitives.lean +++ b/backends/lean/Base/Primitives.lean @@ -4,6 +4,8 @@ import Init.Data.List.Basic import Mathlib.Tactic.RunCmd import Mathlib.Tactic.Linarith +namespace Primitives + -------------------- -- ASSERT COMMAND --Std. -------------------- @@ -46,6 +48,7 @@ open Error inductive Result (α : Type u) where | ret (v: α): Result α | fail (e: Error): Result α + | div deriving Repr, BEq open Result @@ -53,20 +56,28 @@ open Result instance Result_Inhabited (α : Type u) : Inhabited (Result α) := Inhabited.mk (fail panic) +instance Result_Nonempty (α : Type u) : Nonempty (Result α) := + Nonempty.intro div + /- HELPERS -/ def ret? {α: Type} (r: Result α): Bool := match r with - | Result.ret _ => true - | Result.fail _ => false + | ret _ => true + | fail _ | div => false + +def div? {α: Type} (r: Result α): Bool := + match r with + | div => true + | ret _ | fail _ => false def massert (b:Bool) : Result Unit := - if b then .ret () else fail assertionFailure + if b then ret () else fail assertionFailure def eval_global {α: Type} (x: Result α) (_: ret? x): α := match x with - | Result.fail _ => by contradiction - | Result.ret x => x + | fail _ | div => by contradiction + | ret x => x /- DO-DSL SUPPORT -/ @@ -74,6 +85,7 @@ def bind (x: Result α) (f: α -> Result β) : Result β := match x with | ret v => f v | fail v => fail v + | div => div -- Allows using Result in do-blocks instance : Bind Result where @@ -92,8 +104,9 @@ instance : Pure Result where def Result.attach {α: Type} (o : Result α): Result { x : α // o = ret x } := match o with - | .ret x => .ret ⟨x, rfl⟩ - | .fail e => .fail e + | ret x => ret ⟨x, rfl⟩ + | fail e => fail e + | div => div macro "let" e:term " ⟵ " f:term : doElem => `(doElem| let ⟨$e, h⟩ ← Result.attach $f) @@ -648,3 +661,5 @@ def vec_index_mut_back (α : Type u) (v: Vec α) (i: Usize) (x: α): Result (Vec /-- Aeneas-translated function -- useful to reduce non-recursive definitions. Use with `simp [ aeneas ]` -/ register_simp_attr aeneas + +end Primitives -- cgit v1.2.3 From ccc97b46c166a45255096d3fec2444c90f7c5aaa Mon Sep 17 00:00:00 2001 From: Son Ho Date: Wed, 14 Jun 2023 11:24:58 +0200 Subject: Make minor modifications --- backends/lean/Base/Primitives.lean | 1 + 1 file changed, 1 insertion(+) (limited to 'backends/lean/Base/Primitives.lean') diff --git a/backends/lean/Base/Primitives.lean b/backends/lean/Base/Primitives.lean index 85e088fc..6b922143 100644 --- a/backends/lean/Base/Primitives.lean +++ b/backends/lean/Base/Primitives.lean @@ -526,6 +526,7 @@ instance (ty : ScalarTy) : DecidableEq (Scalar ty) := def Scalar.toInt {ty} (n : Scalar ty) : Int := n.val -- Tactic to prove that integers are in bounds +-- TODO: use this: https://leanprover.zulipchat.com/#narrow/stream/270676-lean4/topic/instance.20with.20tactic.20autoparam syntax "intlit" : tactic macro_rules -- cgit v1.2.3 From 04cefd3b4f3d2c11cfc3542a5ad6fae31dae4796 Mon Sep 17 00:00:00 2001 From: Son Ho Date: Sun, 18 Jun 2023 19:09:19 +0200 Subject: Make minor modifications --- backends/lean/Base/Primitives.lean | 5 ++--- 1 file changed, 2 insertions(+), 3 deletions(-) (limited to 'backends/lean/Base/Primitives.lean') diff --git a/backends/lean/Base/Primitives.lean b/backends/lean/Base/Primitives.lean index 6b922143..d6cc0bad 100644 --- a/backends/lean/Base/Primitives.lean +++ b/backends/lean/Base/Primitives.lean @@ -20,9 +20,8 @@ def assertImpl : CommandElab := fun (_stx: Syntax) => do runTermElabM (fun _ => do let r ← evalTerm Bool (mkConst ``Bool) _stx[1] if not r then - logInfo "Assertion failed for: " - logInfo _stx[1] - logError "Expression reduced to false" + logInfo ("Assertion failed for:\n" ++ _stx[1]) + throwError ("Expression reduced to false:\n" ++ _stx[1]) pure ()) #eval 2 == 2 -- cgit v1.2.3 From a2670f4d097075c23b9affceb8ed8498b73c4b8c Mon Sep 17 00:00:00 2001 From: Son Ho Date: Mon, 19 Jun 2023 18:52:29 +0200 Subject: Cleanup Diverge.lean --- backends/lean/Base/Primitives.lean | 13 +++++++++++++ 1 file changed, 13 insertions(+) (limited to 'backends/lean/Base/Primitives.lean') diff --git a/backends/lean/Base/Primitives.lean b/backends/lean/Base/Primitives.lean index d6cc0bad..1185a07d 100644 --- a/backends/lean/Base/Primitives.lean +++ b/backends/lean/Base/Primitives.lean @@ -94,6 +94,10 @@ instance : Bind Result where instance : Pure Result where pure := fun x => ret x +@[simp] theorem bind_ret (x : α) (f : α → Result β) : bind (.ret x) f = f x := by simp [bind] +@[simp] theorem bind_fail (x : Error) (f : α → Result β) : bind (.fail x) f = .fail x := by simp [bind] +@[simp] theorem bind_div (f : α → Result β) : bind .div f = .div := by simp [bind] + /- CUSTOM-DSL SUPPORT -/ -- Let-binding the Result of a monadic operation is oftentimes not sufficient, @@ -124,6 +128,15 @@ macro "let" e:term " <-- " f:term : doElem => let r: { x: Nat // x = 0 } := ⟨ y, by assumption ⟩ .ret r +@[simp] theorem bind_tc_ret (x : α) (f : α → Result β) : + (do let y ← .ret x; f y) = f x := by simp [Bind.bind, bind] + +@[simp] theorem bind_tc_fail (x : Error) (f : α → Result β) : + (do let y ← fail x; f y) = fail x := by simp [Bind.bind, bind] + +@[simp] theorem bind_tc_div (f : α → Result β) : + (do let y ← div; f y) = div := by simp [Bind.bind, bind] + ---------------------- -- MACHINE INTEGERS -- ---------------------- -- cgit v1.2.3 From 4fd17e4bb91eb46d4704643dfbfbbf0874837b07 Mon Sep 17 00:00:00 2001 From: Son Ho Date: Tue, 4 Jul 2023 12:49:37 +0200 Subject: Make Diverge use Primitives --- backends/lean/Base/Primitives.lean | 25 ++++--------------------- 1 file changed, 4 insertions(+), 21 deletions(-) (limited to 'backends/lean/Base/Primitives.lean') diff --git a/backends/lean/Base/Primitives.lean b/backends/lean/Base/Primitives.lean index 1185a07d..117f76a2 100644 --- a/backends/lean/Base/Primitives.lean +++ b/backends/lean/Base/Primitives.lean @@ -60,12 +60,12 @@ instance Result_Nonempty (α : Type u) : Nonempty (Result α) := /- HELPERS -/ -def ret? {α: Type} (r: Result α): Bool := +def ret? {α: Type u} (r: Result α): Bool := match r with | ret _ => true | fail _ | div => false -def div? {α: Type} (r: Result α): Bool := +def div? {α: Type u} (r: Result α): Bool := match r with | div => true | ret _ | fail _ => false @@ -73,14 +73,14 @@ def div? {α: Type} (r: Result α): Bool := def massert (b:Bool) : Result Unit := if b then ret () else fail assertionFailure -def eval_global {α: Type} (x: Result α) (_: ret? x): α := +def eval_global {α: Type u} (x: Result α) (_: ret? x): α := match x with | fail _ | div => by contradiction | ret x => x /- DO-DSL SUPPORT -/ -def bind (x: Result α) (f: α -> Result β) : Result β := +def bind {α : Type u} {β : Type v} (x: Result α) (f: α -> Result β) : Result β := match x with | ret v => f v | fail v => fail v @@ -111,23 +111,6 @@ def Result.attach {α: Type} (o : Result α): Result { x : α // o = ret x } := | fail e => fail e | div => div -macro "let" e:term " ⟵ " f:term : doElem => - `(doElem| let ⟨$e, h⟩ ← Result.attach $f) - --- TODO: any way to factorize both definitions? -macro "let" e:term " <-- " f:term : doElem => - `(doElem| let ⟨$e, h⟩ ← Result.attach $f) - --- We call the hypothesis `h`, in effect making it unavailable to the user --- (because too much shadowing). But in practice, once can use the French single --- quote notation (input with f< and f>), where `‹ h ›` finds a suitable --- hypothesis in the context, this is equivalent to `have x: h := by assumption in x` -#eval do - let y <-- .ret (0: Nat) - let _: y = 0 := by cases ‹ ret 0 = ret y › ; decide - let r: { x: Nat // x = 0 } := ⟨ y, by assumption ⟩ - .ret r - @[simp] theorem bind_tc_ret (x : α) (f : α → Result β) : (do let y ← .ret x; f y) = f x := by simp [Bind.bind, bind] -- cgit v1.2.3 From 5ca36bfc50083a01af2b7ae5f75993a520757ef5 Mon Sep 17 00:00:00 2001 From: Son Ho Date: Wed, 5 Jul 2023 15:17:58 +0200 Subject: Simplify the names used in Primitives.lean --- backends/lean/Base/Primitives.lean | 37 +++++++++++++++++++++---------------- 1 file changed, 21 insertions(+), 16 deletions(-) (limited to 'backends/lean/Base/Primitives.lean') diff --git a/backends/lean/Base/Primitives.lean b/backends/lean/Base/Primitives.lean index 117f76a2..808c1461 100644 --- a/backends/lean/Base/Primitives.lean +++ b/backends/lean/Base/Primitives.lean @@ -564,15 +564,17 @@ macro_rules def Vec (α : Type u) := { l : List α // List.length l ≤ Usize.max } -def vec_new (α : Type u): Vec α := ⟨ [], by apply Scalar.cMax_suffices .Usize; simp ⟩ +def Vec.new (α : Type u): Vec α := ⟨ [], by apply Scalar.cMax_suffices .Usize; simp ⟩ -def vec_len (α : Type u) (v : Vec α) : Usize := +def Vec.len (α : Type u) (v : Vec α) : Usize := let ⟨ v, l ⟩ := v Usize.ofIntCore (List.length v) (by simp [Scalar.min]) l -def vec_push_fwd (α : Type u) (_ : Vec α) (_ : α) : Unit := () +-- This shouldn't be used +def Vec.push_fwd (α : Type u) (_ : Vec α) (_ : α) : Unit := () -def vec_push_back (α : Type u) (v : Vec α) (x : α) : Result (Vec α) +-- This is actually the backward function +def Vec.push (α : Type u) (v : Vec α) (x : α) : Result (Vec α) := let nlen := List.length v.val + 1 if h : nlen ≤ U32.max || nlen ≤ Usize.max then @@ -588,13 +590,15 @@ def vec_push_back (α : Type u) (v : Vec α) (x : α) : Result (Vec α) else fail maximumSizeExceeded -def vec_insert_fwd (α : Type u) (v: Vec α) (i: Usize) (_: α): Result Unit := +-- This shouldn't be used +def Vec.insert_fwd (α : Type u) (v: Vec α) (i: Usize) (_: α): Result Unit := if i.val < List.length v.val then .ret () else .fail arrayOutOfBounds -def vec_insert_back (α : Type u) (v: Vec α) (i: Usize) (x: α): Result (Vec α) := +-- This is actually the backward function +def Vec.insert (α : Type u) (v: Vec α) (i: Usize) (x: α): Result (Vec α) := if i.val < List.length v.val then -- TODO: maybe we should redefine a list library which uses integers -- (instead of natural numbers) @@ -607,7 +611,7 @@ def vec_insert_back (α : Type u) (v: Vec α) (i: Usize) (x: α): Result (Vec α else .fail arrayOutOfBounds -def vec_index_to_fin {α : Type u} {v: Vec α} {i: Usize} (h : i.val < List.length v.val) : +def Vec.index_to_fin {α : Type u} {v: Vec α} {i: Usize} (h : i.val < List.length v.val) : Fin (List.length v.val) := let j := i.val.toNat let h: j < List.length v.val := by @@ -616,29 +620,30 @@ def vec_index_to_fin {α : Type u} {v: Vec α} {i: Usize} (h : i.val < List.leng assumption ⟨j, h⟩ -def vec_index_fwd (α : Type u) (v: Vec α) (i: Usize): Result α := +def Vec.index (α : Type u) (v: Vec α) (i: Usize): Result α := if h: i.val < List.length v.val then - let i := vec_index_to_fin h + let i := Vec.index_to_fin h .ret (List.get v.val i) else .fail arrayOutOfBounds -def vec_index_back (α : Type u) (v: Vec α) (i: Usize) (_: α): Result Unit := +-- This shouldn't be used +def Vec.index_back (α : Type u) (v: Vec α) (i: Usize) (_: α): Result Unit := if i.val < List.length v.val then .ret () else .fail arrayOutOfBounds -def vec_index_mut_fwd (α : Type u) (v: Vec α) (i: Usize): Result α := +def Vec.index_mut (α : Type u) (v: Vec α) (i: Usize): Result α := if h: i.val < List.length v.val then - let i := vec_index_to_fin h + let i := Vec.index_to_fin h .ret (List.get v.val i) else .fail arrayOutOfBounds -def vec_index_mut_back (α : Type u) (v: Vec α) (i: Usize) (x: α): Result (Vec α) := +def Vec.index_mut_back (α : Type u) (v: Vec α) (i: Usize) (x: α): Result (Vec α) := if h: i.val < List.length v.val then - let i := vec_index_to_fin h + let i := Vec.index_to_fin h .ret ⟨ List.set v.val i x, by have h: List.length v.val ≤ Usize.max := v.property simp [*] at * @@ -651,8 +656,8 @@ def vec_index_mut_back (α : Type u) (v: Vec α) (i: Usize) (x: α): Result (Vec -- MISC -- ---------- -@[simp] def mem_replace_fwd (a : Type) (x : a) (_ : a) : a := x -@[simp] def mem_replace_back (a : Type) (_ : a) (y : a) : a := y +@[simp] def mem.replace_fwd (a : Type) (x : a) (_ : a) : a := x +@[simp] def mem.replace_back (a : Type) (_ : a) (y : a) : a := y /-- Aeneas-translated function -- useful to reduce non-recursive definitions. Use with `simp [ aeneas ]` -/ -- cgit v1.2.3 From 7c95800cefc87fad894f8bf855cfc047e713b3a7 Mon Sep 17 00:00:00 2001 From: Son Ho Date: Thu, 6 Jul 2023 12:20:28 +0200 Subject: Improve the generated comments --- backends/lean/Base/Primitives.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) (limited to 'backends/lean/Base/Primitives.lean') diff --git a/backends/lean/Base/Primitives.lean b/backends/lean/Base/Primitives.lean index 808c1461..14f5971e 100644 --- a/backends/lean/Base/Primitives.lean +++ b/backends/lean/Base/Primitives.lean @@ -656,7 +656,7 @@ def Vec.index_mut_back (α : Type u) (v: Vec α) (i: Usize) (x: α): Result (Vec -- MISC -- ---------- -@[simp] def mem.replace_fwd (a : Type) (x : a) (_ : a) : a := x +@[simp] def mem.replace (a : Type) (x : a) (_ : a) : a := x @[simp] def mem.replace_back (a : Type) (_ : a) (y : a) : a := y /-- Aeneas-translated function -- useful to reduce non-recursive definitions. -- cgit v1.2.3 From a18d899a2c2b9bdd36f4a5a4b70472c12a835a96 Mon Sep 17 00:00:00 2001 From: Son Ho Date: Wed, 12 Jul 2023 14:34:55 +0200 Subject: Finish a first version of the progress tactic --- backends/lean/Base/Primitives.lean | 89 ++++++++++++++++++++++++++++---------- 1 file changed, 65 insertions(+), 24 deletions(-) (limited to 'backends/lean/Base/Primitives.lean') diff --git a/backends/lean/Base/Primitives.lean b/backends/lean/Base/Primitives.lean index 14f5971e..6210688d 100644 --- a/backends/lean/Base/Primitives.lean +++ b/backends/lean/Base/Primitives.lean @@ -175,27 +175,28 @@ open System.Platform.getNumBits @[simp] def U128.min : Int := 0 @[simp] def U128.max : Int := HPow.hPow 2 128 - 1 -#assert (I8.min == -128) -#assert (I8.max == 127) -#assert (I16.min == -32768) -#assert (I16.max == 32767) -#assert (I32.min == -2147483648) -#assert (I32.max == 2147483647) -#assert (I64.min == -9223372036854775808) -#assert (I64.max == 9223372036854775807) -#assert (I128.min == -170141183460469231731687303715884105728) -#assert (I128.max == 170141183460469231731687303715884105727) -#assert (U8.min == 0) -#assert (U8.max == 255) -#assert (U16.min == 0) -#assert (U16.max == 65535) -#assert (U32.min == 0) -#assert (U32.max == 4294967295) -#assert (U64.min == 0) -#assert (U64.max == 18446744073709551615) -#assert (U128.min == 0) -#assert (U128.max == 340282366920938463463374607431768211455) - +-- The normalized bounds +@[simp] def I8.norm_min := -128 +@[simp] def I8.norm_max := 127 +@[simp] def I16.norm_min := -32768 +@[simp] def I16.norm_max := 32767 +@[simp] def I32.norm_min := -2147483648 +@[simp] def I32.norm_max := 2147483647 +@[simp] def I64.norm_min := -9223372036854775808 +@[simp] def I64.norm_max := 9223372036854775807 +@[simp] def I128.norm_min := -170141183460469231731687303715884105728 +@[simp] def I128.norm_max := 170141183460469231731687303715884105727 +@[simp] def U8.norm_min := 0 +@[simp] def U8.norm_max := 255 +@[simp] def U16.norm_min := 0 +@[simp] def U16.norm_max := 65535 +@[simp] def U32.norm_min := 0 +@[simp] def U32.norm_max := 4294967295 +@[simp] def U64.norm_min := 0 +@[simp] def U64.norm_max := 18446744073709551615 +@[simp] def U128.norm_min := 0 +@[simp] def U128.norm_max := 340282366920938463463374607431768211455 + inductive ScalarTy := | Isize | I8 @@ -240,6 +241,46 @@ def Scalar.max (ty : ScalarTy) : Int := | .U64 => U64.max | .U128 => U128.max +@[simp] def Scalar.norm_min (ty : ScalarTy) : Int := + match ty with + -- We can't normalize the bounds for isize/usize + | .Isize => Isize.min + | .Usize => Usize.min + -- + | .I8 => I8.norm_min + | .I16 => I16.norm_min + | .I32 => I32.norm_min + | .I64 => I64.norm_min + | .I128 => I128.norm_min + | .U8 => U8.norm_min + | .U16 => U16.norm_min + | .U32 => U32.norm_min + | .U64 => U64.norm_min + | .U128 => U128.norm_min + +@[simp] def Scalar.norm_max (ty : ScalarTy) : Int := + match ty with + -- We can't normalize the bounds for isize/usize + | .Isize => Isize.max + | .Usize => Usize.max + -- + | .I8 => I8.norm_max + | .I16 => I16.norm_max + | .I32 => I32.norm_max + | .I64 => I64.norm_max + | .I128 => I128.norm_max + | .U8 => U8.norm_max + | .U16 => U16.norm_max + | .U32 => U32.norm_max + | .U64 => U64.norm_max + | .U128 => U128.norm_max + +def Scalar.norm_min_eq (ty : ScalarTy) : Scalar.min ty = Scalar.norm_min ty := by + cases ty <;> rfl + +def Scalar.norm_max_eq (ty : ScalarTy) : Scalar.max ty = Scalar.norm_max ty := by + cases ty <;> rfl + -- "Conservative" bounds -- We use those because we can't compare to the isize bounds (which can't -- reduce at compile-time). Whenever we perform an arithmetic operation like @@ -249,13 +290,13 @@ def Scalar.max (ty : ScalarTy) : Int := -- type-checking time. def Scalar.cMin (ty : ScalarTy) : Int := match ty with - | .Isize => I32.min + | .Isize => Scalar.min .I32 | _ => Scalar.min ty def Scalar.cMax (ty : ScalarTy) : Int := match ty with - | .Isize => I32.max - | .Usize => U32.max + | .Isize => Scalar.max .I32 + | .Usize => Scalar.max .U32 | _ => Scalar.max ty theorem Scalar.cMin_bound ty : Scalar.min ty ≤ Scalar.cMin ty := by -- cgit v1.2.3 From 59e4a06480b5365f48dc68de80f44841f94094ed Mon Sep 17 00:00:00 2001 From: Son Ho Date: Wed, 12 Jul 2023 15:41:29 +0200 Subject: Improve the handling of arithmetic bounds --- backends/lean/Base/Primitives.lean | 228 +++++++++++++++++++------------------ 1 file changed, 117 insertions(+), 111 deletions(-) (limited to 'backends/lean/Base/Primitives.lean') diff --git a/backends/lean/Base/Primitives.lean b/backends/lean/Base/Primitives.lean index 6210688d..37abdede 100644 --- a/backends/lean/Base/Primitives.lean +++ b/backends/lean/Base/Primitives.lean @@ -149,54 +149,78 @@ open System.Platform.getNumBits @[simp] def size_num_bits : Nat := (System.Platform.getNumBits ()).val -- Remark: Lean seems to use < for the comparisons with the upper bounds by convention. --- We keep the F* convention for now. -@[simp] def Isize.min : Int := - (HPow.hPow 2 (size_num_bits - 1)) -@[simp] def Isize.max : Int := (HPow.hPow 2 (size_num_bits - 1)) - 1 -@[simp] def I8.min : Int := - (HPow.hPow 2 7) -@[simp] def I8.max : Int := HPow.hPow 2 7 - 1 -@[simp] def I16.min : Int := - (HPow.hPow 2 15) -@[simp] def I16.max : Int := HPow.hPow 2 15 - 1 -@[simp] def I32.min : Int := -(HPow.hPow 2 31) -@[simp] def I32.max : Int := HPow.hPow 2 31 - 1 -@[simp] def I64.min : Int := -(HPow.hPow 2 63) -@[simp] def I64.max : Int := HPow.hPow 2 63 - 1 -@[simp] def I128.min : Int := -(HPow.hPow 2 127) -@[simp] def I128.max : Int := HPow.hPow 2 127 - 1 -@[simp] def Usize.min : Int := 0 -@[simp] def Usize.max : Int := HPow.hPow 2 size_num_bits - 1 -@[simp] def U8.min : Int := 0 -@[simp] def U8.max : Int := HPow.hPow 2 8 - 1 -@[simp] def U16.min : Int := 0 -@[simp] def U16.max : Int := HPow.hPow 2 16 - 1 -@[simp] def U32.min : Int := 0 -@[simp] def U32.max : Int := HPow.hPow 2 32 - 1 -@[simp] def U64.min : Int := 0 -@[simp] def U64.max : Int := HPow.hPow 2 64 - 1 -@[simp] def U128.min : Int := 0 -@[simp] def U128.max : Int := HPow.hPow 2 128 - 1 - --- The normalized bounds -@[simp] def I8.norm_min := -128 -@[simp] def I8.norm_max := 127 -@[simp] def I16.norm_min := -32768 -@[simp] def I16.norm_max := 32767 -@[simp] def I32.norm_min := -2147483648 -@[simp] def I32.norm_max := 2147483647 -@[simp] def I64.norm_min := -9223372036854775808 -@[simp] def I64.norm_max := 9223372036854775807 -@[simp] def I128.norm_min := -170141183460469231731687303715884105728 -@[simp] def I128.norm_max := 170141183460469231731687303715884105727 -@[simp] def U8.norm_min := 0 -@[simp] def U8.norm_max := 255 -@[simp] def U16.norm_min := 0 -@[simp] def U16.norm_max := 65535 -@[simp] def U32.norm_min := 0 -@[simp] def U32.norm_max := 4294967295 -@[simp] def U64.norm_min := 0 -@[simp] def U64.norm_max := 18446744073709551615 -@[simp] def U128.norm_min := 0 -@[simp] def U128.norm_max := 340282366920938463463374607431768211455 - + +-- The "structured" bounds +def Isize.smin : Int := - (HPow.hPow 2 (size_num_bits - 1)) +def Isize.smax : Int := (HPow.hPow 2 (size_num_bits - 1)) - 1 +def I8.smin : Int := - (HPow.hPow 2 7) +def I8.smax : Int := HPow.hPow 2 7 - 1 +def I16.smin : Int := - (HPow.hPow 2 15) +def I16.smax : Int := HPow.hPow 2 15 - 1 +def I32.smin : Int := -(HPow.hPow 2 31) +def I32.smax : Int := HPow.hPow 2 31 - 1 +def I64.smin : Int := -(HPow.hPow 2 63) +def I64.smax : Int := HPow.hPow 2 63 - 1 +def I128.smin : Int := -(HPow.hPow 2 127) +def I128.smax : Int := HPow.hPow 2 127 - 1 +def Usize.smin : Int := 0 +def Usize.smax : Int := HPow.hPow 2 size_num_bits - 1 +def U8.smin : Int := 0 +def U8.smax : Int := HPow.hPow 2 8 - 1 +def U16.smin : Int := 0 +def U16.smax : Int := HPow.hPow 2 16 - 1 +def U32.smin : Int := 0 +def U32.smax : Int := HPow.hPow 2 32 - 1 +def U64.smin : Int := 0 +def U64.smax : Int := HPow.hPow 2 64 - 1 +def U128.smin : Int := 0 +def U128.smax : Int := HPow.hPow 2 128 - 1 + +-- The "normalized" bounds, that we use in practice +def I8.min := -128 +def I8.max := 127 +def I16.min := -32768 +def I16.max := 32767 +def I32.min := -2147483648 +def I32.max := 2147483647 +def I64.min := -9223372036854775808 +def I64.max := 9223372036854775807 +def I128.min := -170141183460469231731687303715884105728 +def I128.max := 170141183460469231731687303715884105727 +@[simp] def U8.min := 0 +def U8.max := 255 +@[simp] def U16.min := 0 +def U16.max := 65535 +@[simp] def U32.min := 0 +def U32.max := 4294967295 +@[simp] def U64.min := 0 +def U64.max := 18446744073709551615 +@[simp] def U128.min := 0 +def U128.max := 340282366920938463463374607431768211455 +@[simp] def Usize.min := 0 + +def Isize.refined_min : { n:Int // n = I32.min ∨ n = I64.min } := + ⟨ Isize.smin, by + simp [Isize.smin] + cases System.Platform.numBits_eq <;> + unfold System.Platform.numBits at * <;> simp [*] ⟩ + +def Isize.refined_max : { n:Int // n = I32.max ∨ n = I64.max } := + ⟨ Isize.smax, by + simp [Isize.smax] + cases System.Platform.numBits_eq <;> + unfold System.Platform.numBits at * <;> simp [*] ⟩ + +def Usize.refined_max : { n:Int // n = U32.max ∨ n = U64.max } := + ⟨ Usize.smax, by + simp [Usize.smax] + cases System.Platform.numBits_eq <;> + unfold System.Platform.numBits at * <;> simp [*] ⟩ + +def Isize.min := Isize.refined_min.val +def Isize.max := Isize.refined_max.val +def Usize.max := Usize.refined_max.val + inductive ScalarTy := | Isize | I8 @@ -211,6 +235,36 @@ inductive ScalarTy := | U64 | U128 +def Scalar.smin (ty : ScalarTy) : Int := + match ty with + | .Isize => Isize.smin + | .I8 => I8.smin + | .I16 => I16.smin + | .I32 => I32.smin + | .I64 => I64.smin + | .I128 => I128.smin + | .Usize => Usize.smin + | .U8 => U8.smin + | .U16 => U16.smin + | .U32 => U32.smin + | .U64 => U64.smin + | .U128 => U128.smin + +def Scalar.smax (ty : ScalarTy) : Int := + match ty with + | .Isize => Isize.smax + | .I8 => I8.smax + | .I16 => I16.smax + | .I32 => I32.smax + | .I64 => I64.smax + | .I128 => I128.smax + | .Usize => Usize.smax + | .U8 => U8.smax + | .U16 => U16.smax + | .U32 => U32.smax + | .U64 => U64.smax + | .U128 => U128.smax + def Scalar.min (ty : ScalarTy) : Int := match ty with | .Isize => Isize.min @@ -241,44 +295,10 @@ def Scalar.max (ty : ScalarTy) : Int := | .U64 => U64.max | .U128 => U128.max -@[simp] def Scalar.norm_min (ty : ScalarTy) : Int := - match ty with - -- We can't normalize the bounds for isize/usize - | .Isize => Isize.min - | .Usize => Usize.min - -- - | .I8 => I8.norm_min - | .I16 => I16.norm_min - | .I32 => I32.norm_min - | .I64 => I64.norm_min - | .I128 => I128.norm_min - | .U8 => U8.norm_min - | .U16 => U16.norm_min - | .U32 => U32.norm_min - | .U64 => U64.norm_min - | .U128 => U128.norm_min - -@[simp] def Scalar.norm_max (ty : ScalarTy) : Int := - match ty with - -- We can't normalize the bounds for isize/usize - | .Isize => Isize.max - | .Usize => Usize.max - -- - | .I8 => I8.norm_max - | .I16 => I16.norm_max - | .I32 => I32.norm_max - | .I64 => I64.norm_max - | .I128 => I128.norm_max - | .U8 => U8.norm_max - | .U16 => U16.norm_max - | .U32 => U32.norm_max - | .U64 => U64.norm_max - | .U128 => U128.norm_max - -def Scalar.norm_min_eq (ty : ScalarTy) : Scalar.min ty = Scalar.norm_min ty := by +def Scalar.smin_eq (ty : ScalarTy) : Scalar.min ty = Scalar.smin ty := by cases ty <;> rfl -def Scalar.norm_max_eq (ty : ScalarTy) : Scalar.max ty = Scalar.norm_max ty := by +def Scalar.smax_eq (ty : ScalarTy) : Scalar.max ty = Scalar.smax ty := by cases ty <;> rfl -- "Conservative" bounds @@ -301,30 +321,22 @@ def Scalar.cMax (ty : ScalarTy) : Int := theorem Scalar.cMin_bound ty : Scalar.min ty ≤ Scalar.cMin ty := by cases ty <;> simp [Scalar.min, Scalar.max, Scalar.cMin, Scalar.cMax] at * - cases System.Platform.numBits_eq <;> - unfold System.Platform.numBits at * <;> - simp [*] + have h := Isize.refined_min.property + cases h <;> simp [*, Isize.min] theorem Scalar.cMax_bound ty : Scalar.cMax ty ≤ Scalar.max ty := by - cases ty <;> simp [Scalar.min, Scalar.max, Scalar.cMin, Scalar.cMax] at * <;> - cases System.Platform.numBits_eq <;> - unfold System.Platform.numBits at * <;> - simp [*] + cases ty <;> simp [Scalar.min, Scalar.max, Scalar.cMin, Scalar.cMax] at * + . have h := Isize.refined_max.property + cases h <;> simp [*, Isize.max] + . have h := Usize.refined_max.property + cases h <;> simp [*, Usize.max] theorem Scalar.cMin_suffices ty (h : Scalar.cMin ty ≤ x) : Scalar.min ty ≤ x := by - cases ty <;> simp [Scalar.min, Scalar.max, Scalar.cMin, Scalar.cMax] at * <;> - cases System.Platform.numBits_eq <;> - unfold System.Platform.numBits at * <;> - simp [*] at * - -- TODO: I would have expected terms like `-(1 + 1) ^ 63` to be simplified + have := Scalar.cMin_bound ty linarith theorem Scalar.cMax_suffices ty (h : x ≤ Scalar.cMax ty) : x ≤ Scalar.max ty := by - cases ty <;> simp [Scalar.min, Scalar.max, Scalar.cMin, Scalar.cMax] at * <;> - cases System.Platform.numBits_eq <;> - unfold System.Platform.numBits at * <;> - simp [*] at * <;> - -- TODO: I would have expected terms like `-(1 + 1) ^ 63` to be simplified + have := Scalar.cMax_bound ty linarith structure Scalar (ty : ScalarTy) where @@ -609,7 +621,7 @@ def Vec.new (α : Type u): Vec α := ⟨ [], by apply Scalar.cMax_suffices .Usiz def Vec.len (α : Type u) (v : Vec α) : Usize := let ⟨ v, l ⟩ := v - Usize.ofIntCore (List.length v) (by simp [Scalar.min]) l + Usize.ofIntCore (List.length v) (by simp [Scalar.min, Usize.min]) l -- This shouldn't be used def Vec.push_fwd (α : Type u) (_ : Vec α) (_ : α) : Unit := () @@ -620,13 +632,9 @@ def Vec.push (α : Type u) (v : Vec α) (x : α) : Result (Vec α) let nlen := List.length v.val + 1 if h : nlen ≤ U32.max || nlen ≤ Usize.max then have h : nlen ≤ Usize.max := by - simp at * - cases System.Platform.numBits_eq <;> - unfold System.Platform.numBits at * <;> - simp [*] at * <;> - try assumption - cases h <;> - linarith + simp [Usize.max] at * + have hm := Usize.refined_max.property + cases h <;> cases hm <;> simp [U32.max, U64.max] at * <;> try linarith return ⟨ List.concat v.val x, by simp at *; assumption ⟩ else fail maximumSizeExceeded @@ -647,7 +655,6 @@ def Vec.insert (α : Type u) (v: Vec α) (i: Usize) (x: α): Result (Vec α) := .ret ⟨ List.set v.val i x, by have h: List.length v.val ≤ Usize.max := v.property simp [*] at * - assumption ⟩ else .fail arrayOutOfBounds @@ -688,7 +695,6 @@ def Vec.index_mut_back (α : Type u) (v: Vec α) (i: Usize) (x: α): Result (Vec .ret ⟨ List.set v.val i x, by have h: List.length v.val ≤ Usize.max := v.property simp [*] at * - assumption ⟩ else .fail arrayOutOfBounds -- cgit v1.2.3 From e010c10fb9a1e2d88b52a4f6b4a0865448276013 Mon Sep 17 00:00:00 2001 From: Son Ho Date: Wed, 12 Jul 2023 15:58:38 +0200 Subject: Make the `by inlit` implicit --- backends/lean/Base/Primitives.lean | 15 +++++++-------- 1 file changed, 7 insertions(+), 8 deletions(-) (limited to 'backends/lean/Base/Primitives.lean') diff --git a/backends/lean/Base/Primitives.lean b/backends/lean/Base/Primitives.lean index 37abdede..0506f4c0 100644 --- a/backends/lean/Base/Primitives.lean +++ b/backends/lean/Base/Primitives.lean @@ -356,8 +356,14 @@ def Scalar.ofIntCore {ty : ScalarTy} (x : Int) (hmin : Scalar.min ty ≤ x) (hmax : x ≤ Scalar.max ty) : Scalar ty := { val := x, hmin := hmin, hmax := hmax } +-- Tactic to prove that integers are in bounds +-- TODO: use this: https://leanprover.zulipchat.com/#narrow/stream/270676-lean4/topic/instance.20with.20tactic.20autoparam +syntax "intlit" : tactic +macro_rules + | `(tactic| intlit) => `(tactic| apply Scalar.bound_suffices; decide) + def Scalar.ofInt {ty : ScalarTy} (x : Int) - (h : Scalar.min ty ≤ x ∧ x ≤ Scalar.max ty) : Scalar ty := + (h : Scalar.min ty ≤ x ∧ x ≤ Scalar.max ty := by intlit) : Scalar ty := -- Remark: we initially wrote: -- let ⟨ hmin, hmax ⟩ := h -- Scalar.ofIntCore x hmin hmax @@ -573,13 +579,6 @@ instance (ty : ScalarTy) : DecidableEq (Scalar ty) := def Scalar.toInt {ty} (n : Scalar ty) : Int := n.val --- Tactic to prove that integers are in bounds --- TODO: use this: https://leanprover.zulipchat.com/#narrow/stream/270676-lean4/topic/instance.20with.20tactic.20autoparam -syntax "intlit" : tactic - -macro_rules - | `(tactic| intlit) => `(tactic| apply Scalar.bound_suffices ; decide) - -- -- We now define a type class that subsumes the various machine integer types, so -- -- as to write a concise definition for scalar_cast, rather than exhaustively -- -- enumerating all of the possible pairs. We remark that Rust has sane semantics -- cgit v1.2.3 From d45c6ed9e8049b81170c3e6950043d08006ba9f2 Mon Sep 17 00:00:00 2001 From: Son Ho Date: Mon, 17 Jul 2023 12:14:03 +0200 Subject: Move a definition --- backends/lean/Base/Primitives.lean | 3 +++ 1 file changed, 3 insertions(+) (limited to 'backends/lean/Base/Primitives.lean') diff --git a/backends/lean/Base/Primitives.lean b/backends/lean/Base/Primitives.lean index 0506f4c0..1a0c665d 100644 --- a/backends/lean/Base/Primitives.lean +++ b/backends/lean/Base/Primitives.lean @@ -616,6 +616,9 @@ def Scalar.toInt {ty} (n : Scalar ty) : Int := n.val def Vec (α : Type u) := { l : List α // List.length l ≤ Usize.max } +-- TODO: do we really need it? It should be with Subtype by default +instance Vec.cast (a : Type): Coe (Vec a) (List a) where coe := λ v => v.val + def Vec.new (α : Type u): Vec α := ⟨ [], by apply Scalar.cMax_suffices .Usize; simp ⟩ def Vec.len (α : Type u) (v : Vec α) : Usize := -- cgit v1.2.3 From 3e8060b5501ec83940a4309389a68898df26ebd0 Mon Sep 17 00:00:00 2001 From: Son Ho Date: Mon, 17 Jul 2023 23:37:31 +0200 Subject: Reorganize the Lean backend --- backends/lean/Base/Primitives.lean | 718 +------------------------------------ 1 file changed, 3 insertions(+), 715 deletions(-) (limited to 'backends/lean/Base/Primitives.lean') diff --git a/backends/lean/Base/Primitives.lean b/backends/lean/Base/Primitives.lean index 1a0c665d..91823cb6 100644 --- a/backends/lean/Base/Primitives.lean +++ b/backends/lean/Base/Primitives.lean @@ -1,715 +1,3 @@ -import Lean -import Lean.Meta.Tactic.Simp -import Init.Data.List.Basic -import Mathlib.Tactic.RunCmd -import Mathlib.Tactic.Linarith - -namespace Primitives - --------------------- --- ASSERT COMMAND --Std. --------------------- - -open Lean Elab Command Term Meta - -syntax (name := assert) "#assert" term: command - -@[command_elab assert] -unsafe -def assertImpl : CommandElab := fun (_stx: Syntax) => do - runTermElabM (fun _ => do - let r ← evalTerm Bool (mkConst ``Bool) _stx[1] - if not r then - logInfo ("Assertion failed for:\n" ++ _stx[1]) - throwError ("Expression reduced to false:\n" ++ _stx[1]) - pure ()) - -#eval 2 == 2 -#assert (2 == 2) - -------------- --- PRELUDE -- -------------- - --- Results & monadic combinators - -inductive Error where - | assertionFailure: Error - | integerOverflow: Error - | divisionByZero: Error - | arrayOutOfBounds: Error - | maximumSizeExceeded: Error - | panic: Error -deriving Repr, BEq - -open Error - -inductive Result (α : Type u) where - | ret (v: α): Result α - | fail (e: Error): Result α - | div -deriving Repr, BEq - -open Result - -instance Result_Inhabited (α : Type u) : Inhabited (Result α) := - Inhabited.mk (fail panic) - -instance Result_Nonempty (α : Type u) : Nonempty (Result α) := - Nonempty.intro div - -/- HELPERS -/ - -def ret? {α: Type u} (r: Result α): Bool := - match r with - | ret _ => true - | fail _ | div => false - -def div? {α: Type u} (r: Result α): Bool := - match r with - | div => true - | ret _ | fail _ => false - -def massert (b:Bool) : Result Unit := - if b then ret () else fail assertionFailure - -def eval_global {α: Type u} (x: Result α) (_: ret? x): α := - match x with - | fail _ | div => by contradiction - | ret x => x - -/- DO-DSL SUPPORT -/ - -def bind {α : Type u} {β : Type v} (x: Result α) (f: α -> Result β) : Result β := - match x with - | ret v => f v - | fail v => fail v - | div => div - --- Allows using Result in do-blocks -instance : Bind Result where - bind := bind - --- Allows using return x in do-blocks -instance : Pure Result where - pure := fun x => ret x - -@[simp] theorem bind_ret (x : α) (f : α → Result β) : bind (.ret x) f = f x := by simp [bind] -@[simp] theorem bind_fail (x : Error) (f : α → Result β) : bind (.fail x) f = .fail x := by simp [bind] -@[simp] theorem bind_div (f : α → Result β) : bind .div f = .div := by simp [bind] - -/- CUSTOM-DSL SUPPORT -/ - --- Let-binding the Result of a monadic operation is oftentimes not sufficient, --- because we may need a hypothesis for equational reasoning in the scope. We --- rely on subtype, and a custom let-binding operator, in effect recreating our --- own variant of the do-dsl - -def Result.attach {α: Type} (o : Result α): Result { x : α // o = ret x } := - match o with - | ret x => ret ⟨x, rfl⟩ - | fail e => fail e - | div => div - -@[simp] theorem bind_tc_ret (x : α) (f : α → Result β) : - (do let y ← .ret x; f y) = f x := by simp [Bind.bind, bind] - -@[simp] theorem bind_tc_fail (x : Error) (f : α → Result β) : - (do let y ← fail x; f y) = fail x := by simp [Bind.bind, bind] - -@[simp] theorem bind_tc_div (f : α → Result β) : - (do let y ← div; f y) = div := by simp [Bind.bind, bind] - ----------------------- --- MACHINE INTEGERS -- ----------------------- - --- We redefine our machine integers types. - --- For Isize/Usize, we reuse `getNumBits` from `USize`. You cannot reduce `getNumBits` --- using the simplifier, meaning that proofs do not depend on the compile-time value of --- USize.size. (Lean assumes 32 or 64-bit platforms, and Rust doesn't really support, at --- least officially, 16-bit microcontrollers, so this seems like a fine design decision --- for now.) - --- Note from Chris Bailey: "If there's more than one salient property of your --- definition then the subtyping strategy might get messy, and the property part --- of a subtype is less discoverable by the simplifier or tactics like --- library_search." So, we will not add refinements on the return values of the --- operations defined on Primitives, but will rather rely on custom lemmas to --- invert on possible return values of the primitive operations. - --- Machine integer constants, done via `ofNatCore`, which requires a proof that --- the `Nat` fits within the desired integer type. We provide a custom tactic. - -open System.Platform.getNumBits - --- TODO: is there a way of only importing System.Platform.getNumBits? --- -@[simp] def size_num_bits : Nat := (System.Platform.getNumBits ()).val - --- Remark: Lean seems to use < for the comparisons with the upper bounds by convention. - --- The "structured" bounds -def Isize.smin : Int := - (HPow.hPow 2 (size_num_bits - 1)) -def Isize.smax : Int := (HPow.hPow 2 (size_num_bits - 1)) - 1 -def I8.smin : Int := - (HPow.hPow 2 7) -def I8.smax : Int := HPow.hPow 2 7 - 1 -def I16.smin : Int := - (HPow.hPow 2 15) -def I16.smax : Int := HPow.hPow 2 15 - 1 -def I32.smin : Int := -(HPow.hPow 2 31) -def I32.smax : Int := HPow.hPow 2 31 - 1 -def I64.smin : Int := -(HPow.hPow 2 63) -def I64.smax : Int := HPow.hPow 2 63 - 1 -def I128.smin : Int := -(HPow.hPow 2 127) -def I128.smax : Int := HPow.hPow 2 127 - 1 -def Usize.smin : Int := 0 -def Usize.smax : Int := HPow.hPow 2 size_num_bits - 1 -def U8.smin : Int := 0 -def U8.smax : Int := HPow.hPow 2 8 - 1 -def U16.smin : Int := 0 -def U16.smax : Int := HPow.hPow 2 16 - 1 -def U32.smin : Int := 0 -def U32.smax : Int := HPow.hPow 2 32 - 1 -def U64.smin : Int := 0 -def U64.smax : Int := HPow.hPow 2 64 - 1 -def U128.smin : Int := 0 -def U128.smax : Int := HPow.hPow 2 128 - 1 - --- The "normalized" bounds, that we use in practice -def I8.min := -128 -def I8.max := 127 -def I16.min := -32768 -def I16.max := 32767 -def I32.min := -2147483648 -def I32.max := 2147483647 -def I64.min := -9223372036854775808 -def I64.max := 9223372036854775807 -def I128.min := -170141183460469231731687303715884105728 -def I128.max := 170141183460469231731687303715884105727 -@[simp] def U8.min := 0 -def U8.max := 255 -@[simp] def U16.min := 0 -def U16.max := 65535 -@[simp] def U32.min := 0 -def U32.max := 4294967295 -@[simp] def U64.min := 0 -def U64.max := 18446744073709551615 -@[simp] def U128.min := 0 -def U128.max := 340282366920938463463374607431768211455 -@[simp] def Usize.min := 0 - -def Isize.refined_min : { n:Int // n = I32.min ∨ n = I64.min } := - ⟨ Isize.smin, by - simp [Isize.smin] - cases System.Platform.numBits_eq <;> - unfold System.Platform.numBits at * <;> simp [*] ⟩ - -def Isize.refined_max : { n:Int // n = I32.max ∨ n = I64.max } := - ⟨ Isize.smax, by - simp [Isize.smax] - cases System.Platform.numBits_eq <;> - unfold System.Platform.numBits at * <;> simp [*] ⟩ - -def Usize.refined_max : { n:Int // n = U32.max ∨ n = U64.max } := - ⟨ Usize.smax, by - simp [Usize.smax] - cases System.Platform.numBits_eq <;> - unfold System.Platform.numBits at * <;> simp [*] ⟩ - -def Isize.min := Isize.refined_min.val -def Isize.max := Isize.refined_max.val -def Usize.max := Usize.refined_max.val - -inductive ScalarTy := -| Isize -| I8 -| I16 -| I32 -| I64 -| I128 -| Usize -| U8 -| U16 -| U32 -| U64 -| U128 - -def Scalar.smin (ty : ScalarTy) : Int := - match ty with - | .Isize => Isize.smin - | .I8 => I8.smin - | .I16 => I16.smin - | .I32 => I32.smin - | .I64 => I64.smin - | .I128 => I128.smin - | .Usize => Usize.smin - | .U8 => U8.smin - | .U16 => U16.smin - | .U32 => U32.smin - | .U64 => U64.smin - | .U128 => U128.smin - -def Scalar.smax (ty : ScalarTy) : Int := - match ty with - | .Isize => Isize.smax - | .I8 => I8.smax - | .I16 => I16.smax - | .I32 => I32.smax - | .I64 => I64.smax - | .I128 => I128.smax - | .Usize => Usize.smax - | .U8 => U8.smax - | .U16 => U16.smax - | .U32 => U32.smax - | .U64 => U64.smax - | .U128 => U128.smax - -def Scalar.min (ty : ScalarTy) : Int := - match ty with - | .Isize => Isize.min - | .I8 => I8.min - | .I16 => I16.min - | .I32 => I32.min - | .I64 => I64.min - | .I128 => I128.min - | .Usize => Usize.min - | .U8 => U8.min - | .U16 => U16.min - | .U32 => U32.min - | .U64 => U64.min - | .U128 => U128.min - -def Scalar.max (ty : ScalarTy) : Int := - match ty with - | .Isize => Isize.max - | .I8 => I8.max - | .I16 => I16.max - | .I32 => I32.max - | .I64 => I64.max - | .I128 => I128.max - | .Usize => Usize.max - | .U8 => U8.max - | .U16 => U16.max - | .U32 => U32.max - | .U64 => U64.max - | .U128 => U128.max - -def Scalar.smin_eq (ty : ScalarTy) : Scalar.min ty = Scalar.smin ty := by - cases ty <;> rfl - -def Scalar.smax_eq (ty : ScalarTy) : Scalar.max ty = Scalar.smax ty := by - cases ty <;> rfl - --- "Conservative" bounds --- We use those because we can't compare to the isize bounds (which can't --- reduce at compile-time). Whenever we perform an arithmetic operation like --- addition we need to check that the result is in bounds: we first compare --- to the conservative bounds, which reduce, then compare to the real bounds. --- This is useful for the various #asserts that we want to reduce at --- type-checking time. -def Scalar.cMin (ty : ScalarTy) : Int := - match ty with - | .Isize => Scalar.min .I32 - | _ => Scalar.min ty - -def Scalar.cMax (ty : ScalarTy) : Int := - match ty with - | .Isize => Scalar.max .I32 - | .Usize => Scalar.max .U32 - | _ => Scalar.max ty - -theorem Scalar.cMin_bound ty : Scalar.min ty ≤ Scalar.cMin ty := by - cases ty <;> simp [Scalar.min, Scalar.max, Scalar.cMin, Scalar.cMax] at * - have h := Isize.refined_min.property - cases h <;> simp [*, Isize.min] - -theorem Scalar.cMax_bound ty : Scalar.cMax ty ≤ Scalar.max ty := by - cases ty <;> simp [Scalar.min, Scalar.max, Scalar.cMin, Scalar.cMax] at * - . have h := Isize.refined_max.property - cases h <;> simp [*, Isize.max] - . have h := Usize.refined_max.property - cases h <;> simp [*, Usize.max] - -theorem Scalar.cMin_suffices ty (h : Scalar.cMin ty ≤ x) : Scalar.min ty ≤ x := by - have := Scalar.cMin_bound ty - linarith - -theorem Scalar.cMax_suffices ty (h : x ≤ Scalar.cMax ty) : x ≤ Scalar.max ty := by - have := Scalar.cMax_bound ty - linarith - -structure Scalar (ty : ScalarTy) where - val : Int - hmin : Scalar.min ty ≤ val - hmax : val ≤ Scalar.max ty -deriving Repr - -theorem Scalar.bound_suffices (ty : ScalarTy) (x : Int) : - Scalar.cMin ty ≤ x ∧ x ≤ Scalar.cMax ty -> - Scalar.min ty ≤ x ∧ x ≤ Scalar.max ty - := - λ h => by - apply And.intro <;> have hmin := Scalar.cMin_bound ty <;> have hmax := Scalar.cMax_bound ty <;> linarith - -def Scalar.ofIntCore {ty : ScalarTy} (x : Int) - (hmin : Scalar.min ty ≤ x) (hmax : x ≤ Scalar.max ty) : Scalar ty := - { val := x, hmin := hmin, hmax := hmax } - --- Tactic to prove that integers are in bounds --- TODO: use this: https://leanprover.zulipchat.com/#narrow/stream/270676-lean4/topic/instance.20with.20tactic.20autoparam -syntax "intlit" : tactic -macro_rules - | `(tactic| intlit) => `(tactic| apply Scalar.bound_suffices; decide) - -def Scalar.ofInt {ty : ScalarTy} (x : Int) - (h : Scalar.min ty ≤ x ∧ x ≤ Scalar.max ty := by intlit) : Scalar ty := - -- Remark: we initially wrote: - -- let ⟨ hmin, hmax ⟩ := h - -- Scalar.ofIntCore x hmin hmax - -- We updated to the line below because a similar pattern in `Scalar.tryMk` - -- made reduction block. Both versions seem to work for `Scalar.ofInt`, though. - -- TODO: investigate - Scalar.ofIntCore x h.left h.right - -@[simp] def Scalar.check_bounds (ty : ScalarTy) (x : Int) : Bool := - (Scalar.cMin ty ≤ x || Scalar.min ty ≤ x) ∧ (x ≤ Scalar.cMax ty || x ≤ Scalar.max ty) - -theorem Scalar.check_bounds_prop {ty : ScalarTy} {x : Int} (h: Scalar.check_bounds ty x) : - Scalar.min ty ≤ x ∧ x ≤ Scalar.max ty := by - simp at * - have ⟨ hmin, hmax ⟩ := h - have hbmin := Scalar.cMin_bound ty - have hbmax := Scalar.cMax_bound ty - cases hmin <;> cases hmax <;> apply And.intro <;> linarith - --- Further thoughts: look at what has been done here: --- https://github.com/leanprover-community/mathlib4/blob/master/Mathlib/Data/Fin/Basic.lean --- and --- https://github.com/leanprover-community/mathlib4/blob/master/Mathlib/Data/UInt.lean --- which both contain a fair amount of reasoning already! -def Scalar.tryMk (ty : ScalarTy) (x : Int) : Result (Scalar ty) := - if h:Scalar.check_bounds ty x then - -- If we do: - -- ``` - -- let ⟨ hmin, hmax ⟩ := (Scalar.check_bounds_prop h) - -- Scalar.ofIntCore x hmin hmax - -- ``` - -- then normalization blocks (for instance, some proofs which use reflexivity fail). - -- However, the version below doesn't block reduction (TODO: investigate): - return Scalar.ofInt x (Scalar.check_bounds_prop h) - else fail integerOverflow - -def Scalar.neg {ty : ScalarTy} (x : Scalar ty) : Result (Scalar ty) := Scalar.tryMk ty (- x.val) - -def Scalar.div {ty : ScalarTy} (x : Scalar ty) (y : Scalar ty) : Result (Scalar ty) := - if y.val != 0 then Scalar.tryMk ty (x.val / y.val) else fail divisionByZero - --- Our custom remainder operation, which satisfies the semantics of Rust --- TODO: is there a better way? -def scalar_rem (x y : Int) : Int := - if 0 ≤ x then |x| % |y| - else - (|x| % |y|) - --- Our custom division operation, which satisfies the semantics of Rust --- TODO: is there a better way? -def scalar_div (x y : Int) : Int := - if 0 ≤ x && 0 ≤ y then |x| / |y| - else if 0 ≤ x && y < 0 then - (|x| / |y|) - else if x < 0 && 0 ≤ y then - (|x| / |y|) - else |x| / |y| - --- Checking that the remainder operation is correct -#assert scalar_rem 1 2 = 1 -#assert scalar_rem (-1) 2 = -1 -#assert scalar_rem 1 (-2) = 1 -#assert scalar_rem (-1) (-2) = -1 -#assert scalar_rem 7 3 = (1:Int) -#assert scalar_rem (-7) 3 = -1 -#assert scalar_rem 7 (-3) = 1 -#assert scalar_rem (-7) (-3) = -1 - --- Checking that the division operation is correct -#assert scalar_div 3 2 = 1 -#assert scalar_div (-3) 2 = -1 -#assert scalar_div 3 (-2) = -1 -#assert scalar_div (-3) (-2) = 1 -#assert scalar_div 7 3 = 2 -#assert scalar_div (-7) 3 = -2 -#assert scalar_div 7 (-3) = -2 -#assert scalar_div (-7) (-3) = 2 - -def Scalar.rem {ty : ScalarTy} (x : Scalar ty) (y : Scalar ty) : Result (Scalar ty) := - if y.val != 0 then Scalar.tryMk ty (x.val % y.val) else fail divisionByZero - -def Scalar.add {ty : ScalarTy} (x : Scalar ty) (y : Scalar ty) : Result (Scalar ty) := - Scalar.tryMk ty (x.val + y.val) - -def Scalar.sub {ty : ScalarTy} (x : Scalar ty) (y : Scalar ty) : Result (Scalar ty) := - Scalar.tryMk ty (x.val - y.val) - -def Scalar.mul {ty : ScalarTy} (x : Scalar ty) (y : Scalar ty) : Result (Scalar ty) := - Scalar.tryMk ty (x.val * y.val) - --- TODO: instances of +, -, * etc. for scalars - --- Cast an integer from a [src_ty] to a [tgt_ty] --- TODO: check the semantics of casts in Rust -def Scalar.cast {src_ty : ScalarTy} (tgt_ty : ScalarTy) (x : Scalar src_ty) : Result (Scalar tgt_ty) := - Scalar.tryMk tgt_ty x.val - --- The scalar types --- We declare the definitions as reducible so that Lean can unfold them (useful --- for type class resolution for instance). -@[reducible] def Isize := Scalar .Isize -@[reducible] def I8 := Scalar .I8 -@[reducible] def I16 := Scalar .I16 -@[reducible] def I32 := Scalar .I32 -@[reducible] def I64 := Scalar .I64 -@[reducible] def I128 := Scalar .I128 -@[reducible] def Usize := Scalar .Usize -@[reducible] def U8 := Scalar .U8 -@[reducible] def U16 := Scalar .U16 -@[reducible] def U32 := Scalar .U32 -@[reducible] def U64 := Scalar .U64 -@[reducible] def U128 := Scalar .U128 - --- TODO: below: not sure this is the best way. --- Should we rather overload operations like +, -, etc.? --- Also, it is possible to automate the generation of those definitions --- with macros (but would it be a good idea? It would be less easy to --- read the file, which is not supposed to change a lot) - --- Negation - -/-- -Remark: there is no heterogeneous negation in the Lean prelude: we thus introduce -one here. - -The notation typeclass for heterogeneous addition. -This enables the notation `- a : β` where `a : α`. --/ -class HNeg (α : Type u) (β : outParam (Type v)) where - /-- `- a` computes the negation of `a`. - The meaning of this notation is type-dependent. -/ - hNeg : α → β - -prefix:75 "-" => HNeg.hNeg - -instance : HNeg Isize (Result Isize) where hNeg x := Scalar.neg x -instance : HNeg I8 (Result I8) where hNeg x := Scalar.neg x -instance : HNeg I16 (Result I16) where hNeg x := Scalar.neg x -instance : HNeg I32 (Result I32) where hNeg x := Scalar.neg x -instance : HNeg I64 (Result I64) where hNeg x := Scalar.neg x -instance : HNeg I128 (Result I128) where hNeg x := Scalar.neg x - --- Addition -instance {ty} : HAdd (Scalar ty) (Scalar ty) (Result (Scalar ty)) where - hAdd x y := Scalar.add x y - --- Substraction -instance {ty} : HSub (Scalar ty) (Scalar ty) (Result (Scalar ty)) where - hSub x y := Scalar.sub x y - --- Multiplication -instance {ty} : HMul (Scalar ty) (Scalar ty) (Result (Scalar ty)) where - hMul x y := Scalar.mul x y - --- Division -instance {ty} : HDiv (Scalar ty) (Scalar ty) (Result (Scalar ty)) where - hDiv x y := Scalar.div x y - --- Remainder -instance {ty} : HMod (Scalar ty) (Scalar ty) (Result (Scalar ty)) where - hMod x y := Scalar.rem x y - --- ofIntCore --- TODO: typeclass? -def Isize.ofIntCore := @Scalar.ofIntCore .Isize -def I8.ofIntCore := @Scalar.ofIntCore .I8 -def I16.ofIntCore := @Scalar.ofIntCore .I16 -def I32.ofIntCore := @Scalar.ofIntCore .I32 -def I64.ofIntCore := @Scalar.ofIntCore .I64 -def I128.ofIntCore := @Scalar.ofIntCore .I128 -def Usize.ofIntCore := @Scalar.ofIntCore .Usize -def U8.ofIntCore := @Scalar.ofIntCore .U8 -def U16.ofIntCore := @Scalar.ofIntCore .U16 -def U32.ofIntCore := @Scalar.ofIntCore .U32 -def U64.ofIntCore := @Scalar.ofIntCore .U64 -def U128.ofIntCore := @Scalar.ofIntCore .U128 - --- ofInt --- TODO: typeclass? -def Isize.ofInt := @Scalar.ofInt .Isize -def I8.ofInt := @Scalar.ofInt .I8 -def I16.ofInt := @Scalar.ofInt .I16 -def I32.ofInt := @Scalar.ofInt .I32 -def I64.ofInt := @Scalar.ofInt .I64 -def I128.ofInt := @Scalar.ofInt .I128 -def Usize.ofInt := @Scalar.ofInt .Usize -def U8.ofInt := @Scalar.ofInt .U8 -def U16.ofInt := @Scalar.ofInt .U16 -def U32.ofInt := @Scalar.ofInt .U32 -def U64.ofInt := @Scalar.ofInt .U64 -def U128.ofInt := @Scalar.ofInt .U128 - --- Comparisons -instance {ty} : LT (Scalar ty) where - lt a b := LT.lt a.val b.val - -instance {ty} : LE (Scalar ty) where le a b := LE.le a.val b.val - -instance Scalar.decLt {ty} (a b : Scalar ty) : Decidable (LT.lt a b) := Int.decLt .. -instance Scalar.decLe {ty} (a b : Scalar ty) : Decidable (LE.le a b) := Int.decLe .. - -theorem Scalar.eq_of_val_eq {ty} : ∀ {i j : Scalar ty}, Eq i.val j.val → Eq i j - | ⟨_, _, _⟩, ⟨_, _, _⟩, rfl => rfl - -theorem Scalar.val_eq_of_eq {ty} {i j : Scalar ty} (h : Eq i j) : Eq i.val j.val := - h ▸ rfl - -theorem Scalar.ne_of_val_ne {ty} {i j : Scalar ty} (h : Not (Eq i.val j.val)) : Not (Eq i j) := - fun h' => absurd (val_eq_of_eq h') h - -instance (ty : ScalarTy) : DecidableEq (Scalar ty) := - fun i j => - match decEq i.val j.val with - | isTrue h => isTrue (Scalar.eq_of_val_eq h) - | isFalse h => isFalse (Scalar.ne_of_val_ne h) - -def Scalar.toInt {ty} (n : Scalar ty) : Int := n.val - --- -- We now define a type class that subsumes the various machine integer types, so --- -- as to write a concise definition for scalar_cast, rather than exhaustively --- -- enumerating all of the possible pairs. We remark that Rust has sane semantics --- -- and fails if a cast operation would involve a truncation or modulo. - --- class MachineInteger (t: Type) where --- size: Nat --- val: t -> Fin size --- ofNatCore: (n:Nat) -> LT.lt n size -> t - --- set_option hygiene false in --- run_cmd --- for typeName in [`UInt8, `UInt16, `UInt32, `UInt64, `USize].map Lean.mkIdent do --- Lean.Elab.Command.elabCommand (← `( --- namespace $typeName --- instance: MachineInteger $typeName where --- size := size --- val := val --- ofNatCore := ofNatCore --- end $typeName --- )) - --- -- Aeneas only instantiates the destination type (`src` is implicit). We rely on --- -- Lean to infer `src`. - --- def scalar_cast { src: Type } (dst: Type) [ MachineInteger src ] [ MachineInteger dst ] (x: src): Result dst := --- if h: MachineInteger.val x < MachineInteger.size dst then --- .ret (MachineInteger.ofNatCore (MachineInteger.val x).val h) --- else --- .fail integerOverflow - -------------- --- VECTORS -- -------------- - -def Vec (α : Type u) := { l : List α // List.length l ≤ Usize.max } - --- TODO: do we really need it? It should be with Subtype by default -instance Vec.cast (a : Type): Coe (Vec a) (List a) where coe := λ v => v.val - -def Vec.new (α : Type u): Vec α := ⟨ [], by apply Scalar.cMax_suffices .Usize; simp ⟩ - -def Vec.len (α : Type u) (v : Vec α) : Usize := - let ⟨ v, l ⟩ := v - Usize.ofIntCore (List.length v) (by simp [Scalar.min, Usize.min]) l - --- This shouldn't be used -def Vec.push_fwd (α : Type u) (_ : Vec α) (_ : α) : Unit := () - --- This is actually the backward function -def Vec.push (α : Type u) (v : Vec α) (x : α) : Result (Vec α) - := - let nlen := List.length v.val + 1 - if h : nlen ≤ U32.max || nlen ≤ Usize.max then - have h : nlen ≤ Usize.max := by - simp [Usize.max] at * - have hm := Usize.refined_max.property - cases h <;> cases hm <;> simp [U32.max, U64.max] at * <;> try linarith - return ⟨ List.concat v.val x, by simp at *; assumption ⟩ - else - fail maximumSizeExceeded - --- This shouldn't be used -def Vec.insert_fwd (α : Type u) (v: Vec α) (i: Usize) (_: α): Result Unit := - if i.val < List.length v.val then - .ret () - else - .fail arrayOutOfBounds - --- This is actually the backward function -def Vec.insert (α : Type u) (v: Vec α) (i: Usize) (x: α): Result (Vec α) := - if i.val < List.length v.val then - -- TODO: maybe we should redefine a list library which uses integers - -- (instead of natural numbers) - let i := i.val.toNat - .ret ⟨ List.set v.val i x, by - have h: List.length v.val ≤ Usize.max := v.property - simp [*] at * - ⟩ - else - .fail arrayOutOfBounds - -def Vec.index_to_fin {α : Type u} {v: Vec α} {i: Usize} (h : i.val < List.length v.val) : - Fin (List.length v.val) := - let j := i.val.toNat - let h: j < List.length v.val := by - have heq := @Int.toNat_lt (List.length v.val) i.val i.hmin - apply heq.mpr - assumption - ⟨j, h⟩ - -def Vec.index (α : Type u) (v: Vec α) (i: Usize): Result α := - if h: i.val < List.length v.val then - let i := Vec.index_to_fin h - .ret (List.get v.val i) - else - .fail arrayOutOfBounds - --- This shouldn't be used -def Vec.index_back (α : Type u) (v: Vec α) (i: Usize) (_: α): Result Unit := - if i.val < List.length v.val then - .ret () - else - .fail arrayOutOfBounds - -def Vec.index_mut (α : Type u) (v: Vec α) (i: Usize): Result α := - if h: i.val < List.length v.val then - let i := Vec.index_to_fin h - .ret (List.get v.val i) - else - .fail arrayOutOfBounds - -def Vec.index_mut_back (α : Type u) (v: Vec α) (i: Usize) (x: α): Result (Vec α) := - if h: i.val < List.length v.val then - let i := Vec.index_to_fin h - .ret ⟨ List.set v.val i x, by - have h: List.length v.val ≤ Usize.max := v.property - simp [*] at * - ⟩ - else - .fail arrayOutOfBounds - ----------- --- MISC -- ----------- - -@[simp] def mem.replace (a : Type) (x : a) (_ : a) : a := x -@[simp] def mem.replace_back (a : Type) (_ : a) (y : a) : a := y - -/-- Aeneas-translated function -- useful to reduce non-recursive definitions. - Use with `simp [ aeneas ]` -/ -register_simp_attr aeneas - -end Primitives +import Base.Primitives.Base +import Base.Primitives.Scalar +import Base.Primitives.Vec -- cgit v1.2.3