NatSet.insert

This commit is contained in:
mehbark 2025-04-04 00:40:08 -04:00
parent 8af59af336
commit c9b1f534be
2 changed files with 40 additions and 53 deletions

View file

@ -19,7 +19,7 @@ def list_insert (x : ) : List → List
| [] => [x] | [] => [x]
| a::as => | a::as =>
if x = a if x = a
then as then a::as
else if x < a else if x < a
then x::a::as then x::a::as
else a::(list_insert x as) else a::(list_insert x as)
@ -30,25 +30,25 @@ theorem list_insert_mem : b ∈ list_insert x xs → b = x b ∈ xs := by
split split
next heq => next heq =>
intro hmem intro hmem
right; right simp_all only [List.mem_cons, or_true]
assumption
next hne => next hne =>
split split
· simp · simp
· next hnlt => next hnlt =>
simp [*] simp [*]
intro h intro h
cases h cases h
· right; left; assumption · right; left; assumption
cases ih _ cases ih _
· left; assumption · left; assumption
· right; right; assumption · right; right; assumption
theorem list_insert_sorted (hs : List.Sorted (·<·) xs) : List.Sorted (·<·) (list_insert x xs) := by theorem list_insert_sorted (hs : List.Sorted (·<·) xs)
: List.Sorted (·<·) (list_insert x xs) := by
induction xs with simp [list_insert] induction xs with simp [list_insert]
| cons a as ih => | cons a as ih =>
split split
· exact List.Sorted.of_cons hs · simp_all only [List.sorted_cons, implies_true, and_self]
· split · split
· exact List.Sorted.cons _ hs · exact List.Sorted.cons _ hs
· have : a < x := by · have : a < x := by
@ -65,43 +65,8 @@ theorem list_insert_sorted (hs : List.Sorted (·<·) xs) : List.Sorted (·<·) (
<;> simp_all only [imp_self, List.sorted_cons, and_true, not_lt] <;> simp_all only [imp_self, List.sorted_cons, and_true, not_lt]
· simp_all only [imp_self, List.sorted_cons, and_true, not_lt] · simp_all only [imp_self, List.sorted_cons, and_true, not_lt]
def insert (x : ) : NatSet → NatSet def insert (x : ) (s : NatSet) : NatSet :=
| ⟨[], _⟩ => singleton x ⟨list_insert x s.val, list_insert_sorted s.prop⟩
| ⟨n::ns, h⟩ =>
if heq : x = n
then ⟨n::ns, h⟩
else if hlt : x < n
then ⟨x::n::ns, List.Sorted.cons hlt h⟩
else
let tail : NatSet := ⟨ns, List.Sorted.of_cons h⟩
let ⟨tail, h₂⟩ := tail.insert x
n::tail,
by
have : ∀b ∈ tail, b = x b ∈ ns := by
intro b h
induction tail
· contradiction
· simp [*] at *
sorry
have : ∀b ∈ ns, n < b := by
simp only [List.sorted_cons] at h
exact h.left
simp only [List.sorted_cons]
constructor
· assumption
-- have : n < x := by
-- rw [Nat.lt_iff_le_and_ne]
-- constructor
-- · simp at hlt
-- assumption
-- · intro h
-- exact heq h.symm
sorry
instance : EmptyCollection NatSet where instance : EmptyCollection NatSet where
emptyCollection := empty emptyCollection := empty
@ -113,5 +78,25 @@ instance : Insert NatSet where
insert := insert insert := insert
instance : LawfulSingleton NatSet where instance : LawfulSingleton NatSet where
insert_empty_eq := by simp [Insert.insert, EmptyCollection.emptyCollection, empty, insert, insert_empty_eq := by intro; rfl
Singleton.singleton, implies_true]
instance : Membership NatSet where
-- this is still an optimization over typical list membership, but
-- TODO: binary search
mem s n := n ∈ s.val.takeWhile (· ≤ n)
instance (n : ) (s : NatSet) : Decidable (n ∈ s) :=
List.instDecidableMemOfLawfulBEq n (s.val.takeWhile (· ≤ n))
-- TODO: (tricky!)
theorem mem_insert (x : ) (s : NatSet) : x ∈ s.insert x := by
sorry
#eval 3 ∈ ({1, 2, 3} : NatSet)
#eval 3 ∈ ({1, 2} : NatSet)
#eval 3 ∈ ({} : NatSet)
#eval ({1, 2} : NatSet) == {3, 4}
#eval ({2, 2, 3, 4} : NatSet) == {4, 3, 2}
#eval ({2, 2, 3, 4} : NatSet)
end NatSet

View file

@ -1 +1,3 @@
# invariant # invariant
basic examples of formally enforcing invariants in data structures, inspired by [`Batteries.dlist`](https://leanprover-community.github.io/mathlib4_docs/Batteries/Data/DList/Basic.html#Batteries.DList)'s `invariant` field.