pairwise attempt

This commit is contained in:
mehbark 2025-04-03 17:42:08 -04:00
commit 592a4b45e6
5 changed files with 70 additions and 0 deletions

47
Invariant.lean Normal file
View file

@ -0,0 +1,47 @@
-- sorted set of natural numbers
structure NatSet where
items : List Nat
invariant : List.Pairwise (· < ·) items
namespace NatSet
-- TODO: binary search
instance : Membership Nat NatSet where
mem l x := x ∈ l.items
instance (x : Nat) (l : NatSet) : Decidable (x ∈ l) :=
List.instDecidableMemOfLawfulBEq x l.items
def empty : NatSet := ⟨[], by simp⟩
def singleton (x : Nat) : NatSet :=
⟨[x], by simp⟩
def remove (x : Nat) (l : NatSet) : NatSet where
items := l.items.erase x
invariant := List.Pairwise.erase x l.invariant
def insert (x : Nat) : NatSet → NatSet
| ⟨[], h⟩ => ⟨[x], List.pairwise_singleton ..⟩
| l@⟨a::as, h⟩ =>
if a = x
then l
else if hlt : x < a
then ⟨
x::a::as,
by
apply List.pairwise_cons.mpr
constructor
· intro a' hmem
have := List.Pairwise.iff
sorry
· assumption
else
let ⟨rest, hr⟩ := (NatSet.mk as (List.Pairwise.of_cons h)).insert x
a::rest,
by
simp [*]
sorry,

1
README.md Normal file
View file

@ -0,0 +1 @@
# invariant

15
lake-manifest.json Normal file
View file

@ -0,0 +1,15 @@
{"version": "1.1.0",
"packagesDir": ".lake/packages",
"packages":
[{"url": "https://github.com/leanprover-community/batteries",
"type": "git",
"subDir": null,
"scope": "leanprover-community",
"rev": "ff4cd734811e5ab091cef2acb4d805737670eced",
"name": "batteries",
"manifestFile": "lake-manifest.json",
"inputRev": "main",
"inherited": false,
"configFile": "lakefile.toml"}],
"name": "invariant",
"lakeDir": ".lake"}

6
lakefile.toml Normal file
View file

@ -0,0 +1,6 @@
name = "invariant"
version = "0.1.0"
defaultTargets = ["Invariant"]
[[lean_lib]]
name = "Invariant"

1
lean-toolchain Normal file
View file

@ -0,0 +1 @@
leanprover/lean4:v4.19.0-rc2