From 5e35e3875884ca4be63f253d9e8d7f2fc71b3527 Mon Sep 17 00:00:00 2001 From: Raito Bezarius Date: Thu, 28 Mar 2024 17:59:35 +0100 Subject: refactor: generalize the theory and perform some lifts Move forward the "HSpec" idea, move around files, construct the hierarchy of trees, etc. Signed-off-by: Raito Bezarius --- AvlVerification/Insert.lean | 87 +++++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 87 insertions(+) create mode 100644 AvlVerification/Insert.lean (limited to 'AvlVerification/Insert.lean') diff --git a/AvlVerification/Insert.lean b/AvlVerification/Insert.lean new file mode 100644 index 0000000..870e9c2 --- /dev/null +++ b/AvlVerification/Insert.lean @@ -0,0 +1,87 @@ +import AvlVerification.Tree +import AvlVerification.BinarySearchTree +import AvlVerification.Specifications + +namespace Implementation + +open Primitives +open avl_verification +open Tree (AVLTree) + +variable {T: Type} + +@[pspec] +theorem AVLTreeSet.insert_loop_spec {H: avl_verification.Ord T} + (a: T) (t: Option (AVLNode T)) + [LT T] + (Hbs_search: BST.BST t) + (Hcmp_spec: ∀ a b, ∃ o, H.cmp a b = Result.ret o): + ∃ added t_new, AVLTreeSet.insert_loop T H a t = Result.ret ⟨ added, t_new ⟩ +-- ∧ AVLTree.set t'.2 = insert a (AVLTree.set t) + ∧ BST.BST t_new + := by match t with + | none => + simp [AVLTreeSet.insert_loop, AVLTree.set, setOf, BST.BST] + | some (AVLNode.mk b left right) => + rw [AVLTreeSet.insert_loop] + simp only [] + progress keep Hordering as ⟨ ordering ⟩ + cases ordering + all_goals simp only [] + have Hbs_search_right: BST.BST right := BST.right Hbs_search + have Hbs_search_left: BST.BST left := BST.left Hbs_search + { + progress keep Htree as ⟨ added₁, right₁, Hbst ⟩ + refine' ⟨ added₁, ⟨ some (AVLNode.mk b left right₁), _ ⟩ ⟩ + simp only [true_and] + refine' (BST.BST.some b left right₁ _ _ Hbs_search_left Hbst) + cases' Hbs_search; assumption + induction right with + | none => + simp [AVLTreeSet.insert_loop] at Htree + rw [← Htree.2] + refine' (BST.ForallNode.some _ _ _ BST.ForallNode.none _ BST.ForallNode.none) + sorry + | some node => + -- clef: ∀ x ∈ node.right, b < x + -- de plus: b < a + -- or, right₁ = insert a node.right moralement + -- donc, il suffit de prouver que: ∀ x ∈ insert a node.right, b < x <-> b < a /\ ∀ x ∈ node.right, b < x + sorry + } + { + simp [Hbs_search] + } + { + -- Symmetric case of left. + sorry + } + +@[pspec] +def AVLTreeSet.insert_spec + {H: avl_verification.Ord T} + -- TODO: this can be generalized into `H.cmp` must be an equivalence relation. + -- and insert works no longer on Sets but on set quotiented by this equivalence relation. + (Hcmp_eq: ∀ a b, H.cmp a b = Result.ret Ordering.Equal -> a = b) + (Hcmp_spec: ∀ a b, ∃ o, H.cmp a b = Result.ret o) + (Hbs_search: @BSTree.searchInvariant _ H t) + (a: T) (t: AVLTreeSet T): + ∃ t', t.insert _ H a = Result.ret t' + -- set of values *POST* insertion is {a} \cup set of values of the *PRE* tree. + ∧ AVLTree.setOf t'.2.root = insert a (AVLTree.setOf t.root) + -- it's still a binary search tree. + ∧ @BSTree.searchInvariant _ H t'.2.root + -- TODO(reinforcement): (t'.1 is false <=> a \in AVLTree.setOf t.root) + := by + rw [AVLTreeSet.insert] + progress keep h as ⟨ t', Hset ⟩; simp + rw [Hset] + sorry + + +-- insert is infaillible, right? +-- instance [LT T] (o: avl_verification.Ord T): Insert T (AVLTreeSet T) where +-- insert x tree := (AVLTreeSet.insert T o tree x).map (fun (added, tree) => tree) + +end Implementation + -- cgit v1.2.3 From 22a499f5d4d231bef3193405983f9ade6da116db Mon Sep 17 00:00:00 2001 From: Raito Bezarius Date: Thu, 4 Apr 2024 17:30:46 +0200 Subject: feat: close key theorem for any result on binary search trees Signed-off-by: Raito Bezarius --- AvlVerification/Insert.lean | 51 +++++++++++++++++++++++---------------------- 1 file changed, 26 insertions(+), 25 deletions(-) (limited to 'AvlVerification/Insert.lean') diff --git a/AvlVerification/Insert.lean b/AvlVerification/Insert.lean index 870e9c2..92691b2 100644 --- a/AvlVerification/Insert.lean +++ b/AvlVerification/Insert.lean @@ -11,49 +11,50 @@ open Tree (AVLTree) variable {T: Type} @[pspec] -theorem AVLTreeSet.insert_loop_spec {H: avl_verification.Ord T} - (a: T) (t: Option (AVLNode T)) - [LT T] - (Hbs_search: BST.BST t) - (Hcmp_spec: ∀ a b, ∃ o, H.cmp a b = Result.ret o): +theorem AVLTreeSet.insert_loop_spec_local (p: T -> Prop) + (Hcmp_spec: ∀ a b, ∃ o, H.cmp a b = Result.ret o) + (a: T) (t: Option (AVLNode T)): ∃ added t_new, AVLTreeSet.insert_loop T H a t = Result.ret ⟨ added, t_new ⟩ -- ∧ AVLTree.set t'.2 = insert a (AVLTree.set t) - ∧ BST.BST t_new + ∧ (BST.ForallNode p t -> p a -> BST.ForallNode p t_new) := by match t with | none => - simp [AVLTreeSet.insert_loop, AVLTree.set, setOf, BST.BST] + simp [AVLTreeSet.insert_loop, AVLTree.set, setOf] + intros _ Hpa + constructor; all_goals try simp [BST.ForallNode.none] + exact Hpa | some (AVLNode.mk b left right) => rw [AVLTreeSet.insert_loop] simp only [] progress keep Hordering as ⟨ ordering ⟩ cases ordering all_goals simp only [] - have Hbs_search_right: BST.BST right := BST.right Hbs_search - have Hbs_search_left: BST.BST left := BST.left Hbs_search { - progress keep Htree as ⟨ added₁, right₁, Hbst ⟩ + progress keep Htree as ⟨ added₁, right₁, Hnode ⟩ refine' ⟨ added₁, ⟨ some (AVLNode.mk b left right₁), _ ⟩ ⟩ simp only [true_and] - refine' (BST.BST.some b left right₁ _ _ Hbs_search_left Hbst) - cases' Hbs_search; assumption - induction right with - | none => - simp [AVLTreeSet.insert_loop] at Htree - rw [← Htree.2] - refine' (BST.ForallNode.some _ _ _ BST.ForallNode.none _ BST.ForallNode.none) - sorry - | some node => - -- clef: ∀ x ∈ node.right, b < x - -- de plus: b < a - -- or, right₁ = insert a node.right moralement - -- donc, il suffit de prouver que: ∀ x ∈ insert a node.right, b < x <-> b < a /\ ∀ x ∈ node.right, b < x - sorry + intros Hptree Hpa + constructor + exact Hptree.left + exact Hptree.label + exact Hnode Hptree.right Hpa } { - simp [Hbs_search] + simp; tauto } { + -- TODO: investigate wlog. -- Symmetric case of left. + progress keep Htree as ⟨ added₁, left₁, Hnode ⟩ + refine' ⟨ added₁, ⟨ some (AVLNode.mk b left₁ right), _ ⟩ ⟩ + simp only [true_and] + intros Hptree Hpa + constructor + exact Hnode Hptree.left Hpa + exact Hptree.label + exact Hptree.right + } + sorry } -- cgit v1.2.3 From 50be416e22a39eaf59f7edba81b919e1f114a0ae Mon Sep 17 00:00:00 2001 From: Raito Bezarius Date: Thu, 4 Apr 2024 17:31:19 +0200 Subject: feat: close the BST proof modulo unbundling Signed-off-by: Raito Bezarius --- AvlVerification/Insert.lean | 51 ++++++++++++++++++++++++++++++++++++++++++++- 1 file changed, 50 insertions(+), 1 deletion(-) (limited to 'AvlVerification/Insert.lean') diff --git a/AvlVerification/Insert.lean b/AvlVerification/Insert.lean index 92691b2..3c2b8b3 100644 --- a/AvlVerification/Insert.lean +++ b/AvlVerification/Insert.lean @@ -7,8 +7,16 @@ namespace Implementation open Primitives open avl_verification open Tree (AVLTree) +open Specifications (OrdSpec ordSpecOfInfaillible ordOfOrdSpec ltOfRustOrder) -variable {T: Type} +-- TODO: OSpec should be proven. +variable (T: Type) (H: avl_verification.Ord T) (Ospec: @OrdSpec T H) + +instance rustOrder {U: Type} {O: avl_verification.Ord U} (OSpec: OrdSpec O): _root_.Ord U := ordOfOrdSpec O OSpec +-- Why the TC fails if I don't specify the previous instance explicitly? +instance rustLt {U: Type} {O: avl_verification.Ord U} (OSpec: OrdSpec O): LT U := @ltOfOrd _ (ordOfOrdSpec O OSpec) + +instance rustLt' : LT T := rustLt Ospec @[pspec] theorem AVLTreeSet.insert_loop_spec_local (p: T -> Prop) @@ -55,6 +63,47 @@ theorem AVLTreeSet.insert_loop_spec_local (p: T -> Prop) exact Hptree.right } +@[pspec] +lemma AVLTreeSet.insert_loop_spec_global + (a: T) (t: Option (AVLNode T)) + : + haveI : LT T := (rustLt Ospec) + BST.BST t -> ∃ added t_new, AVLTreeSet.insert_loop T H a t = Result.ret ⟨ added, t_new ⟩ + ∧ BST.BST t_new := by + intro Hbst + letI instOrd : _root_.Ord T := (rustOrder Ospec) + letI instLt : LT T := (rustLt Ospec) + match t with + | none => simp [AVLTreeSet.insert_loop] + | some (AVLNode.mk b left right) => + rw [AVLTreeSet.insert_loop] + simp only [] + have := Ospec.infallible; simp at this + -- TODO: unbundle `this`... + have : ∀ a b, ∃ o, H.cmp a b = .ret o := sorry + progress keep Hordering as ⟨ ordering ⟩ + cases ordering + all_goals simp only [] + { + have ⟨ added₂, right₂, ⟨ H_result, H_bst ⟩ ⟩ := AVLTreeSet.insert_loop_spec_global a right (BST.right Hbst) + progress keep Htree with AVLTreeSet.insert_loop_spec_local as ⟨ added₁, right₁, Hnode ⟩ + exact (fun x => b < x) + rewrite [Htree] at H_result; simp at H_result + refine' ⟨ added₁, ⟨ some (AVLNode.mk b left right₁), _ ⟩ ⟩ + simp only [true_and] + cases' Hbst with _ _ _ H_forall_left H_forall_right H_bst_left H_bst_right + constructor + exact H_forall_left + apply Hnode; exact H_forall_right + exact (ltOfRustOrder H b a Hordering) + exact H_bst_left + convert H_bst + exact H_result.2 + } + { + simp; tauto + } + { sorry } -- cgit v1.2.3 From 3577c7dc9d3013d401c45a7628b0ff4b6fd0ec67 Mon Sep 17 00:00:00 2001 From: Raito Bezarius Date: Fri, 12 Apr 2024 19:11:16 +0200 Subject: feat: upgrade and close all proofs except Preorder on U32 Signed-off-by: Raito Bezarius --- AvlVerification/Insert.lean | 95 +++++++++++++++++++++++++++++++-------------- 1 file changed, 65 insertions(+), 30 deletions(-) (limited to 'AvlVerification/Insert.lean') diff --git a/AvlVerification/Insert.lean b/AvlVerification/Insert.lean index 3c2b8b3..baf3441 100644 --- a/AvlVerification/Insert.lean +++ b/AvlVerification/Insert.lean @@ -7,9 +7,47 @@ namespace Implementation open Primitives open avl_verification open Tree (AVLTree) -open Specifications (OrdSpec ordSpecOfInfaillible ordOfOrdSpec ltOfRustOrder) +open Specifications (OrdSpec ordSpecOfTotalityAndDuality ordOfOrdSpec ltOfRustOrder gtOfRustOrder) + +example: OrdSpec OrdU32 := ordSpecOfTotalityAndDuality _ + (by + -- Totality + intro a b + unfold Ord.cmp + unfold OrdU32 + unfold OrdU32.cmp + if hlt : a < b then + use .Less + simp [hlt] + else + if heq: a = b + then + use .Equal + simp [hlt] + rw [heq] + -- TODO: simp [hlt, heq] breaks everything??? + else + use .Greater + simp [hlt, heq] + ) (by + -- Duality + intro a b Hgt + if hlt : b < a then + unfold Ord.cmp + unfold OrdU32 + unfold OrdU32.cmp + simp [hlt] + else + unfold Ord.cmp at Hgt + unfold OrdU32 at Hgt + unfold OrdU32.cmp at Hgt + have hnlt : ¬ (a < b) := sorry + have hneq : ¬ (a = b) := sorry + exfalso + apply hlt + -- I need a Preorder on U32 now. + sorry) --- TODO: OSpec should be proven. variable (T: Type) (H: avl_verification.Ord T) (Ospec: @OrdSpec T H) instance rustOrder {U: Type} {O: avl_verification.Ord U} (OSpec: OrdSpec O): _root_.Ord U := ordOfOrdSpec O OSpec @@ -20,9 +58,9 @@ instance rustLt' : LT T := rustLt Ospec @[pspec] theorem AVLTreeSet.insert_loop_spec_local (p: T -> Prop) - (Hcmp_spec: ∀ a b, ∃ o, H.cmp a b = Result.ret o) + (Hcmp_spec: ∀ a b, ∃ o, H.cmp a b = Result.ok o) (a: T) (t: Option (AVLNode T)): - ∃ added t_new, AVLTreeSet.insert_loop T H a t = Result.ret ⟨ added, t_new ⟩ + ∃ added t_new, AVLTreeSet.insert_loop T H a t = Result.ok ⟨ added, t_new ⟩ -- ∧ AVLTree.set t'.2 = insert a (AVLTree.set t) ∧ (BST.ForallNode p t -> p a -> BST.ForallNode p t_new) := by match t with @@ -68,8 +106,8 @@ lemma AVLTreeSet.insert_loop_spec_global (a: T) (t: Option (AVLNode T)) : haveI : LT T := (rustLt Ospec) - BST.BST t -> ∃ added t_new, AVLTreeSet.insert_loop T H a t = Result.ret ⟨ added, t_new ⟩ - ∧ BST.BST t_new := by + BST.Invariant t -> ∃ added t_new, AVLTreeSet.insert_loop T H a t = Result.ok ⟨ added, t_new ⟩ + ∧ BST.Invariant t_new := by intro Hbst letI instOrd : _root_.Ord T := (rustOrder Ospec) letI instLt : LT T := (rustLt Ospec) @@ -78,9 +116,7 @@ lemma AVLTreeSet.insert_loop_spec_global | some (AVLNode.mk b left right) => rw [AVLTreeSet.insert_loop] simp only [] - have := Ospec.infallible; simp at this - -- TODO: unbundle `this`... - have : ∀ a b, ∃ o, H.cmp a b = .ret o := sorry + have : ∀ a b, ∃ o, H.cmp a b = .ok o := Ospec.infallible progress keep Hordering as ⟨ ordering ⟩ cases ordering all_goals simp only [] @@ -104,34 +140,33 @@ lemma AVLTreeSet.insert_loop_spec_global simp; tauto } { - sorry + have ⟨ added₂, left₂, ⟨ H_result, H_bst ⟩ ⟩ := AVLTreeSet.insert_loop_spec_global a left (BST.left Hbst) + progress keep Htree with AVLTreeSet.insert_loop_spec_local as ⟨ added₁, left₁, Hnode ⟩ + exact (fun x => x < b) + rewrite [Htree] at H_result; simp at H_result + refine' ⟨ added₁, ⟨ some (AVLNode.mk b left₁ right), _ ⟩ ⟩ + simp only [true_and] + cases' Hbst with _ _ _ H_forall_left H_forall_right H_bst_left H_bst_right + constructor + apply Hnode; exact H_forall_left + exact (gtOfRustOrder H b a Hordering) + exact H_forall_right + convert H_bst + exact H_result.2 + exact H_bst_right } @[pspec] def AVLTreeSet.insert_spec - {H: avl_verification.Ord T} - -- TODO: this can be generalized into `H.cmp` must be an equivalence relation. - -- and insert works no longer on Sets but on set quotiented by this equivalence relation. - (Hcmp_eq: ∀ a b, H.cmp a b = Result.ret Ordering.Equal -> a = b) - (Hcmp_spec: ∀ a b, ∃ o, H.cmp a b = Result.ret o) - (Hbs_search: @BSTree.searchInvariant _ H t) (a: T) (t: AVLTreeSet T): - ∃ t', t.insert _ H a = Result.ret t' - -- set of values *POST* insertion is {a} \cup set of values of the *PRE* tree. - ∧ AVLTree.setOf t'.2.root = insert a (AVLTree.setOf t.root) + haveI : LT T := (rustLt Ospec) + BST.Invariant t.root -> (∃ t' added,t.insert _ H a = Result.ok (added, t') -- it's still a binary search tree. - ∧ @BSTree.searchInvariant _ H t'.2.root - -- TODO(reinforcement): (t'.1 is false <=> a \in AVLTree.setOf t.root) + ∧ BST.Invariant t'.root) := by - rw [AVLTreeSet.insert] - progress keep h as ⟨ t', Hset ⟩; simp - rw [Hset] - sorry - - --- insert is infaillible, right? --- instance [LT T] (o: avl_verification.Ord T): Insert T (AVLTreeSet T) where --- insert x tree := (AVLTreeSet.insert T o tree x).map (fun (added, tree) => tree) + rw [AVLTreeSet.insert]; intro Hbst + progress keep h as ⟨ t', Hset ⟩; + simp; assumption end Implementation -- cgit v1.2.3