From dee74ca1f90acb076289286f6f69df65e63604ce Mon Sep 17 00:00:00 2001 From: Jonathan Protzenko Date: Mon, 23 Jan 2023 18:43:45 -0800 Subject: Write a tactic to discharge integer literal proof obligations --- backends/lean/primitives.lean | 27 +++++++++++++++++++++++---- 1 file changed, 23 insertions(+), 4 deletions(-) (limited to 'backends/lean/primitives.lean') diff --git a/backends/lean/primitives.lean b/backends/lean/primitives.lean index b68df5f0..d86c0423 100644 --- a/backends/lean/primitives.lean +++ b/backends/lean/primitives.lean @@ -83,12 +83,13 @@ def massert (b:Bool) : result Unit := -- and -- https://github.com/leanprover-community/mathlib4/blob/master/Mathlib/Data/UInt.lean -- which both contain a fair amount of reasoning already! -def USize.checked_sub (n: USize) (m: Nat): result USize := +def USize.checked_sub (n: USize) (m: USize): result USize := -- NOTE: the test USize.toNat n - m >= 0 seems to always succeed? - if USize.toNat n >= m then + if n >= m then let n' := USize.toNat n - let r := USize.ofNatCore (n' - m) (by - have h: n' - m <= n' := by + let m' := USize.toNat n + let r := USize.ofNatCore (n' - m') (by + have h: n' - m' <= n' := by apply Nat.sub_le_of_le_add case h => rewrite [ Nat.add_comm ]; apply Nat.le_add_left apply Nat.lt_of_le_of_lt h @@ -102,6 +103,24 @@ def USize.checked_sub (n: USize) (m: Nat): result USize := def USize.checked_mul (n: USize) (m: USize): result USize := sorry def USize.checked_div (n: USize) (m: USize): result USize := sorry +-- One needs to perform a little bit of reasoning in order to successfully +-- inject constants into USize, so we provide a general-purpose macro + +syntax "intlit" : tactic + +macro_rules + | `(tactic| intlit) => `(tactic| + match USize.size, usize_size_eq with + | _, Or.inl rfl => decide + | _, Or.inr rfl => decide) + +-- This is how the macro is expected to be used +#eval USize.ofNatCore 0 (by intlit) + +-- Also works for other integer types (at the expense of a needless disjunction) +#eval UInt32.ofNatCore 0 (by intlit) + +-- Test behavior... #eval USize.checked_sub 10 20 #eval USize.checked_sub 20 10 -- NOTE: compare with concrete behavior here, which I do not think we want -- cgit v1.2.3