diff options
author | Ryan Lahfa | 2024-04-23 14:26:27 +0200 |
---|---|---|
committer | GitHub | 2024-04-23 14:26:27 +0200 |
commit | d3ea366adcd71d0ef15ffbc5d35ca998aa53f19a (patch) | |
tree | d9da70f7564ea73ceacf880b78473c89f617bba7 /Verification/Find.lean | |
parent | 0f49a61cc33bddf2cc69bc8915b95c915dc5f987 (diff) | |
parent | b650710ad3f8c14b713bdf52f684f472115dce2f (diff) |
Merge pull request #3 from RaitoBezarius/bst-find
feat: `find` and `insert` reinforced proofs
Diffstat (limited to 'Verification/Find.lean')
-rw-r--r-- | Verification/Find.lean | 40 |
1 files changed, 40 insertions, 0 deletions
diff --git a/Verification/Find.lean b/Verification/Find.lean new file mode 100644 index 0000000..764a685 --- /dev/null +++ b/Verification/Find.lean @@ -0,0 +1,40 @@ +import Verification.Tree +import Verification.BinarySearchTree +import Verification.Specifications +import AvlVerification + +namespace Implementation + +open Primitives +open avl_verification +open Tree (AVLTree AVLTree.set) +open Specifications (OrdSpecLinearOrderEq infallible ltOfRustOrder gtOfRustOrder) + +variable (T: Type) (H: avl_verification.Ord T) [DecidableEq T] [LinearOrder T] (Ospec: OrdSpecLinearOrderEq H) + +@[pspec] +def AVLTreeSet.find_loop_spec + (a: T) (t: Option (AVLNode T)): + BST.Invariant t -> a ∈ AVLTree.set t -> AVLTreeSet.find_loop _ H a t = Result.ok true := fun Hbst Hmem => by + match t with + | none => trivial + | some (AVLNode.mk b left right) => + rw [AVLTreeSet.find_loop] + dsimp only + have : ∀ a b, ∃ o, H.cmp a b = .ok o := infallible H + progress keep Hordering as ⟨ ordering ⟩ + cases ordering + all_goals dsimp only + . refine' AVLTreeSet.find_loop_spec a right (BST.right Hbst) (BST.right_pos Hbst Hmem _) + exact ltOfRustOrder _ _ _ Hordering + . refine' AVLTreeSet.find_loop_spec a left (BST.left Hbst) (BST.left_pos Hbst Hmem _) + exact gtOfRustOrder _ _ _ Hordering + +def AVLTreeSet.find_spec + (a: T) (t: AVLTreeSet T): + BST.Invariant t.root -> a ∈ AVLTree.set t.root -> + t.find _ H a = Result.ok true := fun Hbst Hmem => by + rw [AVLTreeSet.find]; progress + +end Implementation + |