summaryrefslogtreecommitdiff
path: root/backends
diff options
context:
space:
mode:
authorSon Ho2023-07-25 19:06:05 +0200
committerSon Ho2023-07-25 19:06:05 +0200
commit0cc3c78137434d848188eee2a66b1e2cacfd102e (patch)
tree43c9cee6a846f634ecd9f144c3c3f51934c7f81e /backends
parent1854c631a6a7a3f8d45ad18e05547f9d3782c3ee (diff)
Make progress on the proofs of the hashmap
Diffstat (limited to 'backends')
-rw-r--r--backends/lean/Base/Arith/Int.lean1
-rw-r--r--backends/lean/Base/IList/IList.lean41
-rw-r--r--backends/lean/Base/Primitives/Base.lean2
-rw-r--r--backends/lean/Base/Primitives/Vec.lean20
-rw-r--r--backends/lean/Base/Progress/Progress.lean34
-rw-r--r--backends/lean/Base/Utils.lean36
6 files changed, 104 insertions, 30 deletions
diff --git a/backends/lean/Base/Arith/Int.lean b/backends/lean/Base/Arith/Int.lean
index bc0676d8..48a30a49 100644
--- a/backends/lean/Base/Arith/Int.lean
+++ b/backends/lean/Base/Arith/Int.lean
@@ -43,6 +43,7 @@ instance (x y : Int) : IsLinearIntProp (x > y) where
instance (x y : Int) : IsLinearIntProp (x ≤ y) where
instance (x y : Int) : IsLinearIntProp (x ≥ y) where
instance (x y : Int) : IsLinearIntProp (x ≥ y) where
+instance (x y : Int) : IsLinearIntProp (x = y) where
/- It seems we don't need to do any special preprocessing when the *goal*
has the following shape - I guess `linarith` automatically calls `intro` -/
instance (x y : Int) : IsLinearIntProp (¬ x = y) where
diff --git a/backends/lean/Base/IList/IList.lean b/backends/lean/Base/IList/IList.lean
index 2443b1a6..93047a1b 100644
--- a/backends/lean/Base/IList/IList.lean
+++ b/backends/lean/Base/IList/IList.lean
@@ -123,6 +123,10 @@ theorem length_update (ls : List α) (i : Int) (x : α) : (ls.update i x).length
theorem len_update (ls : List α) (i : Int) (x : α) : (ls.update i x).len = ls.len := by
simp [len_eq_length]
+@[simp]
+theorem len_map (ls : List α) (f : α → β) : (ls.map f).len = ls.len := by
+ simp [len_eq_length]
+
theorem left_length_eq_append_eq (l1 l2 l1' l2' : List α) (heq : l1.length = l1'.length) :
l1 ++ l2 = l1' ++ l2' ↔ l1 = l1' ∧ l2 = l2' := by
revert l1'
@@ -210,6 +214,43 @@ theorem index_eq
simp at *
apply index_eq <;> scalar_tac
+theorem update_map_eq {α : Type u} {β : Type v} (ls : List α) (i : Int) (x : α) (f : α → β) :
+ (ls.update i x).map f = (ls.map f).update i (f x) :=
+ match ls with
+ | [] => by simp
+ | hd :: tl =>
+ if h : i = 0 then by simp [*]
+ else
+ have hi := update_map_eq tl (i - 1) x f
+ by simp [*]
+
+theorem len_flatten_update_eq {α : Type u} (ls : List (List α)) (i : Int) (x : List α)
+ (h0 : 0 ≤ i) (h1 : i < ls.len) :
+ (ls.update i x).flatten.len = ls.flatten.len + x.len - (ls.index i).len :=
+ match ls with
+ | [] => by simp at h1; int_tac
+ | hd :: tl => by
+ simp at h1
+ if h : i = 0 then simp [*]; int_tac
+ else
+ have hi := len_flatten_update_eq tl (i - 1) x (by int_tac) (by int_tac)
+ simp [*]
+ int_tac
+
+@[simp]
+theorem index_map_eq {α : Type u} {β : Type v} [Inhabited α] [Inhabited β] (ls : List α) (i : Int) (f : α → β)
+ (h0 : 0 ≤ i) (h1 : i < ls.len) :
+ (ls.map f).index i = f (ls.index i) :=
+ match ls with
+ | [] => by simp at h1; int_tac
+ | hd :: tl =>
+ if h : i = 0 then by
+ simp [*]
+ else
+ have hi := index_map_eq tl (i - 1) f (by int_tac) (by simp at h1; int_tac)
+ by
+ simp [*]
+
def allP {α : Type u} (l : List α) (p: α → Prop) : Prop :=
foldr (fun a r => p a ∧ r) True l
diff --git a/backends/lean/Base/Primitives/Base.lean b/backends/lean/Base/Primitives/Base.lean
index db462c38..7c0fa3bb 100644
--- a/backends/lean/Base/Primitives/Base.lean
+++ b/backends/lean/Base/Primitives/Base.lean
@@ -76,7 +76,7 @@ def eval_global {α: Type u} (x: Result α) (_: ret? x): α :=
/- DO-DSL SUPPORT -/
-def bind {α : Type u} {β : Type v} (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
diff --git a/backends/lean/Base/Primitives/Vec.lean b/backends/lean/Base/Primitives/Vec.lean
index 5a709566..523372bb 100644
--- a/backends/lean/Base/Primitives/Vec.lean
+++ b/backends/lean/Base/Primitives/Vec.lean
@@ -75,10 +75,9 @@ def Vec.insert (α : Type u) (v: Vec α) (i: Usize) (x: α) : Result (Vec α) :=
.fail arrayOutOfBounds
@[pspec]
-theorem Vec.insert_spec {α : Type u} (v: Vec α) (i: Usize) (x: α) :
- i.val < v.length →
+theorem Vec.insert_spec {α : Type u} (v: Vec α) (i: Usize) (x: α)
+ (hbound : i.val < v.length) :
∃ nv, v.insert α i x = ret nv ∧ nv.val = v.val.update i.val x := by
- intro h
simp [insert, *]
def Vec.index (α : Type u) (v: Vec α) (i: Usize) : Result α :=
@@ -87,10 +86,9 @@ def Vec.index (α : Type u) (v: Vec α) (i: Usize) : Result α :=
| some x => ret x
@[pspec]
-theorem Vec.index_spec {α : Type u} [Inhabited α] (v: Vec α) (i: Usize) :
- i.val < v.length →
+theorem Vec.index_spec {α : Type u} [Inhabited α] (v: Vec α) (i: Usize)
+ (hbound : i.val < v.length) :
v.index α i = ret (v.val.index i.val) := by
- intro
simp only [index]
-- TODO: dependent rewrite
have h := List.indexOpt_eq_index v.val i.val (by scalar_tac) (by simp [*])
@@ -109,10 +107,9 @@ def Vec.index_mut (α : Type u) (v: Vec α) (i: Usize) : Result α :=
| some x => ret x
@[pspec]
-theorem Vec.index_mut_spec {α : Type u} [Inhabited α] (v: Vec α) (i: Usize) :
- i.val < v.length →
+theorem Vec.index_mut_spec {α : Type u} [Inhabited α] (v: Vec α) (i: Usize)
+ (hbound : i.val < v.length) :
v.index_mut α i = ret (v.val.index i.val) := by
- intro
simp only [index_mut]
-- TODO: dependent rewrite
have h := List.indexOpt_eq_index v.val i.val (by scalar_tac) (by simp [*])
@@ -129,12 +126,11 @@ def Vec.index_mut_back (α : Type u) (v: Vec α) (i: Usize) (x: α) : Result (Ve
.ret ⟨ v.val.update i.val x, by have := v.property; simp [*] ⟩
@[pspec]
-theorem Vec.index_mut_back_spec {α : Type u} (v: Vec α) (i: Usize) (x : α) :
- i.val < v.length →
+theorem Vec.index_mut_back_spec {α : Type u} (v: Vec α) (i: Usize) (x : α)
+ (hbound : i.val < v.length) :
∃ nv, v.index_mut_back α i x = ret nv ∧
nv.val = v.val.update i.val x
:= by
- intro
simp only [index_mut_back]
have h := List.indexOpt_bounds v.val i.val
split
diff --git a/backends/lean/Base/Progress/Progress.lean b/backends/lean/Base/Progress/Progress.lean
index a2c7764f..4a406bdf 100644
--- a/backends/lean/Base/Progress/Progress.lean
+++ b/backends/lean/Base/Progress/Progress.lean
@@ -1,6 +1,7 @@
import Lean
import Base.Arith
import Base.Progress.Base
+import Base.Primitives -- TODO: remove?
namespace Progress
@@ -41,7 +42,12 @@ def progressWith (fExpr : Expr) (th : TheoremOrLocal)
match th with
| .Theorem thName =>
let thDecl := env.constants.find! thName
- pure thDecl.type
+ -- 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
| .Local asmDecl => pure asmDecl.type
trace[Progress] "Looked up theorem/assumption type: {thTy}"
-- TODO: the tactic fails if we uncomment withNewMCtxDepth
@@ -129,15 +135,16 @@ def progressWith (fExpr : Expr) (th : TheoremOrLocal)
Split the remaining conjunctions by using fresh ids if the user
instructed to fully split the post-condition, otherwise stop -/
if splitPost then
- splitFullConjTac hPost (λ _ => pure .Ok)
+ splitFullConjTac true hPost (λ _ => pure .Ok)
else pure .Ok
| nid :: ids => do
- trace[Progress] "Splitting post: {hPost}"
+ trace[Progress] "Splitting post: {← inferType hPost}"
-- Split
let nid ← do
match nid with
| none => mkFreshUserName `h
| some nid => pure nid
+ trace[Progress] "\n- prevId: {prevId}\n- nid: {nid}\n- remaining ids: {ids}"
if ← isConj (← inferType hPost) then
splitConjTac hPost (some (prevId, nid)) (λ _ nhPost => splitPostWithIds nid nhPost ids)
else return (.Error m!"Too many ids provided ({ids0}) not enough conjuncts to split in the postcondition")
@@ -323,7 +330,7 @@ def evalProgress (args : TSyntax `Progress.progressArgs) : TacticM Unit := do
elab "progress" args:progressArgs : tactic =>
evalProgress args
-/- namespace Test
+namespace Test
open Primitives Result
set_option trace.Progress true
@@ -336,10 +343,25 @@ elab "progress" args:progressArgs : tactic =>
(hmin : Scalar.min ty ≤ x.val + y.val)
(hmax : x.val + y.val ≤ Scalar.max ty) :
∃ z, x + y = ret z ∧ z.val = x.val + y.val := by
--- progress keep as h with Scalar.add_spec as ⟨ z ⟩
progress keep as h as ⟨ x, h1 .. ⟩
simp [*]
-end Test -/
+ example {ty} {x y : Scalar ty}
+ (hmin : Scalar.min ty ≤ x.val + y.val)
+ (hmax : x.val + y.val ≤ Scalar.max ty) :
+ ∃ z, x + y = ret z ∧ z.val = x.val + y.val := by
+ progress keep as h with Scalar.add_spec as ⟨ z ⟩
+ simp [*]
+
+ /- Checking that universe instantiation works: the original spec uses
+ `α : Type u` where u is quantified, while here we use `α : Type 0` -/
+ example {α : Type} (v: Vec α) (i: Usize) (x : α)
+ (hbounds : i.val < v.length) :
+ ∃ nv, v.index_mut_back α i x = ret nv ∧
+ nv.val = v.val.update i.val x := by
+ progress
+ simp [*]
+
+end Test
end Progress
diff --git a/backends/lean/Base/Utils.lean b/backends/lean/Base/Utils.lean
index 66497a49..f6dc45c7 100644
--- a/backends/lean/Base/Utils.lean
+++ b/backends/lean/Base/Utils.lean
@@ -1,6 +1,7 @@
import Lean
import Mathlib.Tactic.Core
import Mathlib.Tactic.LeftRight
+import Base.UtilsBase
/-
Mathlib tactics:
@@ -331,13 +332,13 @@ def assumptionTac : TacticM Unit :=
liftMetaTactic fun mvarId => do mvarId.assumption; pure []
def isConj (e : Expr) : MetaM Bool :=
- e.withApp fun f args => pure (f.isConstOf ``And ∧ args.size = 2)
+ e.consumeMData.withApp fun f args => pure (f.isConstOf ``And ∧ args.size = 2)
-- Return the first conjunct if the expression is a conjunction, or the
-- expression itself otherwise. Also return the second conjunct if it is a
-- conjunction.
def optSplitConj (e : Expr) : MetaM (Expr × Option Expr) := do
- e.withApp fun f args =>
+ e.consumeMData.withApp fun f args =>
if f.isConstOf ``And ∧ args.size = 2 then pure (args.get! 0, some (args.get! 1))
else pure (e, none)
@@ -345,6 +346,7 @@ def optSplitConj (e : Expr) : MetaM (Expr × Option Expr) := do
def splitConjTarget : TacticM Unit := do
withMainContext do
let g ← getMainTarget
+ trace[Utils] "splitConjTarget: goal: {g}"
-- The tactic was initially implemened with `_root_.Lean.MVarId.apply`
-- but it tended to mess the goal by unfolding terms, even when it failed
let (l, r) ← optSplitConj g
@@ -525,18 +527,26 @@ def splitConjTac (h : Expr) (optIds : Option (Name × Name)) (k : Expr → Expr
throwError "Not a conjunction"
-- Tactic to fully split a conjunction
-partial def splitFullConjTacAux [Inhabited α] [Nonempty α] (l : List Expr) (h : Expr) (k : List Expr → TacticM α) : TacticM α := do
+partial def splitFullConjTacAux [Inhabited α] [Nonempty α] (keepCurrentName : Bool) (l : List Expr) (h : Expr) (k : List Expr → TacticM α) : TacticM α := do
try
- splitConjTac h none (λ h1 h2 =>
- splitFullConjTacAux l h1 (λ l1 =>
- splitFullConjTacAux l1 h2 (λ l2 =>
+ let ids ← do
+ if keepCurrentName then do
+ let cur := (← h.fvarId!.getDecl).userName
+ let nid ← mkFreshUserName `h
+ pure (some (cur, nid))
+ else
+ pure none
+ splitConjTac h ids (λ h1 h2 =>
+ splitFullConjTacAux keepCurrentName l h1 (λ l1 =>
+ splitFullConjTacAux keepCurrentName l1 h2 (λ l2 =>
k l2)))
catch _ =>
k (h :: l)
-- Tactic to fully split a conjunction
-def splitFullConjTac [Inhabited α] [Nonempty α] (h : Expr) (k : List Expr → TacticM α) : TacticM α := do
- splitFullConjTacAux [] h (λ l => k l.reverse)
+-- `keepCurrentName`: if `true`, then the first conjunct has the name of the original assumption
+def splitFullConjTac [Inhabited α] [Nonempty α] (keepCurrentName : Bool) (h : Expr) (k : List Expr → TacticM α) : TacticM α := do
+ splitFullConjTacAux keepCurrentName [] h (λ l => k l.reverse)
syntax optAtArgs := ("at" ident)?
def elabOptAtArgs (args : TSyntax `Utils.optAtArgs) : TacticM (Option Expr) := do
@@ -553,17 +563,21 @@ def elabOptAtArgs (args : TSyntax `Utils.optAtArgs) : TacticM (Option Expr) := d
elab "split_conj" args:optAtArgs : tactic => do
withMainContext do
match ← elabOptAtArgs args with
- | some fvar =>
+ | some fvar => do
+ trace[Utils] "split at {fvar}"
splitConjTac fvar none (fun _ _ => pure ())
- | none =>
+ | none => do
+ trace[Utils] "split goal"
splitConjTarget
elab "split_conjs" args:optAtArgs : tactic => do
withMainContext do
match ← elabOptAtArgs args with
| some fvar =>
- splitFullConjTac fvar (fun _ => pure ())
+ trace[Utils] "split at {fvar}"
+ splitFullConjTac false fvar (fun _ => pure ())
| none =>
+ trace[Utils] "split goal"
repeatTac splitConjTarget
elab "split_existsl" " at " n:ident : tactic => do