summaryrefslogtreecommitdiff
path: root/tests/lean/Hashmap/Properties.lean
diff options
context:
space:
mode:
Diffstat (limited to 'tests/lean/Hashmap/Properties.lean')
-rw-r--r--tests/lean/Hashmap/Properties.lean76
1 files changed, 65 insertions, 11 deletions
diff --git a/tests/lean/Hashmap/Properties.lean b/tests/lean/Hashmap/Properties.lean
index baa8f5c8..e065bb36 100644
--- a/tests/lean/Hashmap/Properties.lean
+++ b/tests/lean/Hashmap/Properties.lean
@@ -21,8 +21,9 @@ end List
namespace HashMap
-@[pspec]
-theorem insert_in_list_spec {α : Type} (key: Usize) (value: α) (ls: List α) :
+abbrev Core.List := _root_.List
+
+theorem insert_in_list_spec0 {α : Type} (key: Usize) (value: α) (ls: List α) :
∃ b,
insert_in_list α key value ls = ret b ∧
(b ↔ List.lookup ls.v key = none)
@@ -35,7 +36,7 @@ theorem insert_in_list_spec {α : Type} (key: Usize) (value: α) (ls: List α) :
rw [insert_in_list_loop]
simp [h]
else
- have ⟨ b, hi ⟩ := insert_in_list_spec key value tl
+ have ⟨ b, hi ⟩ := insert_in_list_spec0 key value tl
by
exists b
simp [insert_in_list, List.lookup]
@@ -45,7 +46,6 @@ theorem insert_in_list_spec {α : Type} (key: Usize) (value: α) (ls: List α) :
exact hi
-- Variation: use progress
-@[pspec]
theorem insert_in_list_spec1 {α : Type} (key: Usize) (value: α) (ls: List α) :
∃ b,
insert_in_list α key value ls = ret b ∧
@@ -70,7 +70,6 @@ theorem insert_in_list_spec1 {α : Type} (key: Usize) (value: α) (ls: List α)
simp [*, List.lookup]
-- Variation: use tactics from the beginning
-@[pspec]
theorem insert_in_list_spec2 {α : Type} (key: Usize) (value: α) (ls: List α) :
∃ b,
insert_in_list α key value ls = ret b ∧
@@ -91,10 +90,10 @@ theorem insert_in_list_spec2 {α : Type} (key: Usize) (value: α) (ls: List α)
progress as ⟨ b h ⟩
simp [*]
-@[pspec]
theorem insert_in_list_back_spec {α : Type} (key: Usize) (value: α) (l0: List α) :
∃ l1,
insert_in_list_back α key value l0 = ret l1 ∧
+ -- We update the binding
List.lookup l1.v key = value ∧
(∀ k, k ≠ key → List.lookup l1.v k = List.lookup l0.v k)
:= match l0 with
@@ -112,11 +111,66 @@ theorem insert_in_list_back_spec {α : Type} (key: Usize) (value: α) (l0: List
simp [insert_in_list_back, List.lookup]
rw [insert_in_list_loop_back]
simp [h, List.lookup]
- have ⟨tl0 , _, _ ⟩ := insert_in_list_back_spec key value tl -- TODO: Use progress
- simp [insert_in_list_back] at *
- simp [List.lookup, *]
- simp (config := {contextual := true}) [*]
-
+ progress keep as heq as ⟨ tl hp1 hp2 ⟩
+ simp [insert_in_list_back] at heq
+ simp (config := {contextual := true}) [*, List.lookup]
+
+def distinct_keys (ls : Core.List (Usize × α)) := ls.pairwise_rel (λ x y => x.fst ≠ y.fst)
+
+def hash_mod_key (k : Usize) (l : Int) : Int :=
+ match hash_key k with
+ | .ret k => k.val % l
+ | _ => 0
+
+def slot_s_inv_hash (l i : Int) (ls : Core.List (Usize × α)) : Prop :=
+ ls.allP (λ (k, _) => hash_mod_key k l = i)
+
+@[simp]
+def slot_s_inv (l i : Int) (ls : Core.List (Usize × α)) : Prop :=
+ distinct_keys ls ∧
+ slot_s_inv_hash l i ls
+
+def slot_t_inv (l i : Int) (s : List α) : Prop := slot_s_inv l i s.v
+
+@[pspec]
+theorem insert_in_list_back_spec1 {α : Type} (l : Int) (key: Usize) (value: α) (l0: List α)
+ (hinv : slot_s_inv_hash l (hash_mod_key key l) l0.v) :
+ ∃ l1,
+ insert_in_list_back α key value l0 = ret l1 ∧
+ -- We update the binding
+ List.lookup l1.v key = value ∧
+ (∀ k, k ≠ key → List.lookup l1.v k = List.lookup l0.v k) ∧
+ -- We preserve part of the key invariant
+ slot_s_inv_hash l (hash_mod_key key l) l1.v
+ := match l0 with
+ | .Nil => by
+ simp [insert_in_list_back, insert_in_list_loop_back, List.lookup, List.v, slot_s_inv_hash]
+ tauto
+ | .Cons k v tl0 =>
+ if h: k = key then
+ by
+ simp [insert_in_list_back, List.lookup]
+ rw [insert_in_list_loop_back]
+ simp [h, List.lookup]
+ constructor
+ . intros; simp [*]
+ . simp [List.v, slot_s_inv_hash] at *
+ simp [*]
+ else
+ by
+ simp [insert_in_list_back, List.lookup]
+ rw [insert_in_list_loop_back]
+ simp [h, List.lookup]
+ have : slot_s_inv_hash l (hash_mod_key key l) (List.v tl0) := by
+ simp_all [List.v, slot_s_inv_hash]
+ progress keep as heq as ⟨ tl1 hp1 hp2 hp3 ⟩
+ simp only [insert_in_list_back] at heq
+ have : slot_s_inv_hash l (hash_mod_key key l) (List.v (List.Cons k v tl1)) := by
+ simp [List.v, slot_s_inv_hash] at *
+ simp [*]
+ simp (config := {contextual := true}) [*, List.lookup]
+
+
end HashMap
end hashmap