Skip to content

Commit

Permalink
Prove properties about delete and nub
Browse files Browse the repository at this point in the history
  • Loading branch information
HeinrichApfelmus committed Jan 29, 2025
1 parent 7bef1a5 commit 2cf0f3d
Showing 1 changed file with 63 additions and 0 deletions.
63 changes: 63 additions & 0 deletions lib/Haskell/Data/List.agda
Original file line number Diff line number Diff line change
Expand Up @@ -38,3 +38,66 @@ sortOn f =
map snd
∘ sortBy (comparing fst)
∘ map (λ x let y = f x in seq y (y , x))

{-----------------------------------------------------------------------------
Properties
------------------------------------------------------------------------------}
--
lemma-neq-trans
: ∀ ⦃ _ : Eq a ⦄ ⦃ _ : IsLawfulEq a ⦄
(x y z : a)
(x == z) ≡ True
(y == z) ≡ False
(x == y) ≡ False
--
lemma-neq-trans x y z eqxz
rewrite equality x z eqxz
rewrite eqSymmetry y z
= λ x x

-- | A deleted item is no longer an element.
--
prop-elem-delete
: ∀ ⦃ _ : Eq a ⦄ ⦃ _ : IsLawfulEq a ⦄
(x y : a) (zs : List a)
elem x (delete y zs)
≡ (if x == y then False else elem x zs)
--
prop-elem-delete x y []
with x == y
... | False = refl
... | True = refl
prop-elem-delete x y (z ∷ zs)
with recurse prop-elem-delete x y zs
with y == z in eqyz
... | True
with x == z in eqxz
... | True
rewrite equality' _ _ (trans (equality x z eqxz) (sym (equality y z eqyz)))
= recurse
... | False
= recurse
prop-elem-delete x y (z ∷ zs)
| False
with x == z in eqxz
... | True
rewrite (lemma-neq-trans x y z eqxz eqyz)
= refl
... | False
= recurse

-- | An item is an element of the 'nub' iff it is
-- an element of the original list.
--
prop-elem-nub
: ∀ ⦃ _ : Eq a ⦄ ⦃ _ : IsLawfulEq a ⦄
(x : a) (ys : List a)
elem x (nub ys)
≡ elem x ys
--
prop-elem-nub x [] = refl
prop-elem-nub x (y ∷ ys)
rewrite prop-elem-delete x y (nub ys)
with x == y
... | True = refl
... | False = prop-elem-nub x ys

0 comments on commit 2cf0f3d

Please sign in to comment.