Skip to content

Commit

Permalink
Merge pull request #1 from leanprover/simp!
Browse files Browse the repository at this point in the history
Use fuel for loop iterations only, simplify final proof with it
  • Loading branch information
nomeata authored Oct 2, 2024
2 parents c43a1d8 + 3b5f21b commit 488e456
Show file tree
Hide file tree
Showing 2 changed files with 30 additions and 30 deletions.
2 changes: 1 addition & 1 deletion Imp.lean
Original file line number Diff line number Diff line change
Expand Up @@ -69,5 +69,5 @@ def test_popcount (x : BitVec 32) : Bool :=


theorem popCount_correct : test_popcount x := by
simp [test_popcount, run, popcount, Expr.eval, Expr.BinOp.apply, Env.get, Env.set, Value, pop_spec, pop_spec.go]
simp! [test_popcount, popcount, pop_spec, Stmt.run]
bv_decide
58 changes: 29 additions & 29 deletions Imp/Stmt/BigStep.lean
Original file line number Diff line number Diff line change
Expand Up @@ -239,35 +239,35 @@ theorem optimize_ok : BigStep σ s σ' → BigStep σ s.optimize σ' := by
assumption

/--
Run a program, with the depth of the recursive calls limited by the `Nat` parameter. Returns `none`
Run a program, with the number of loop iterations limited by the `Nat` parameter. Returns `none`
if the program doesn't terminate fast enough or if some other problem means the result is undefined
(e.g. division by zero).
-/
def run (σ : Env) (s : Stmt) : Nat → Option Env
| 0 => none
| n + 1 =>
match s with
| imp {skip;} =>
some σ
| imp {~s1 ~s2} => do
def run (σ : Env) (s : Stmt) (n : Nat) : Option Env :=
match s with
| imp {skip;} =>
some σ
| imp {~s1 ~s2} => do
let σ' ← run σ s1 n
run σ' s2 n
| imp {~x := ~e;} => do
let v ← e.eval σ
pure (σ.set x v)
| imp {if (~c) {~s1} else {~s2}} => do
let v ← c.eval σ
if v = 0 then
run σ s2 n
else
run σ s1 n
| imp {while (~c) {~s1}} => do
let n + 1 := n
| none
let v ← c.eval σ
if v = 0 then
pure σ
else
let σ' ← run σ s1 n
run σ' s2 n
| imp {~x := ~e;} => do
let v ← e.eval σ
pure (σ.set x v)
| imp {if (~c) {~s1} else {~s2}} => do
let v ← c.eval σ
if v = 0 then
run σ s2 n
else
run σ s1 n
| imp {while (~c) {~s1}} => do
let v ← c.eval σ
if v = 0 then
pure σ
else
let σ' ← run σ s1 n
run σ' (imp {while (~c) {~s1}}) n
run σ' (imp {while (~c) {~s1}}) n

/-- Helper lemma for the next proof -/
theorem bind_eq_some : (Option.bind x f = some y) ↔ (∃ a, x = some a ∧ f a = some y) := by
Expand All @@ -289,18 +289,18 @@ confidence.
theorem run_some_implies_big_step : run σ s n = some σ' → BigStep σ s σ' := by
intro term
induction σ, s, n using run.induct generalizing σ' <;> simp_all [run, bind_eq_some]
case case3 σ n s1 s2 ih1 ih2 =>
case case2 σ n s1 s2 ih1 ih2 =>
let ⟨σ'', run_s1, run_s2⟩ := term
-- Right nesting needed to run the first `trivial` before the second `apply_assumption`, which solves a metavariable
constructor <;> (apply_assumption <;> trivial)
case case4 =>
case case3 =>
let ⟨v, has_val, next_env⟩ := term
subst_eqs
constructor <;> trivial
case case5 ih1 ih2 =>
case case4 ih1 ih2 =>
let ⟨v, has_val, next_env⟩ := term
split at next_env <;> simp_all [BigStep.ifTrue, BigStep.ifFalse]
case case6 ih1 ih2 =>
case case5 ih1 ih2 =>
let ⟨v, has_val, next_env⟩ := term
split at next_env
. simp_all [BigStep.whileFalse]
Expand Down

0 comments on commit 488e456

Please sign in to comment.