Skip to content
New issue

Have a question about this project? Sign up for a free GitHub account to open an issue and contact its maintainers and the community.

By clicking “Sign up for GitHub”, you agree to our terms of service and privacy statement. We’ll occasionally send you account related emails.

Already on GitHub? Sign in to your account

Refactor Algebra.Solver.*Monoid #2407

Merged
merged 33 commits into from
Sep 3, 2024
Merged
Show file tree
Hide file tree
Changes from 9 commits
Commits
Show all changes
33 commits
Select commit Hold shift + click to select a range
1e5eec1
refactor `Solver` infrastructure
jamesmckinna Jun 8, 2024
036f8b5
refactor `Solver` infrastructure
jamesmckinna Jun 9, 2024
8c6342c
tighten imports and some `with`
jamesmckinna Jun 9, 2024
e8ac5fd
tidy up `Solver.Monoid`
jamesmckinna Jun 9, 2024
a495a12
refactor `Algebra.Solver.*Monoid`
jamesmckinna Jun 11, 2024
77f420a
regularise vocabulary: deprecations
jamesmckinna Jun 11, 2024
54a53a8
`variable`s for `Algebra.Solver.Ring`
jamesmckinna Jun 11, 2024
706b3a1
tweaks
jamesmckinna Jun 11, 2024
24d0495
forgot to add `Expression` module
jamesmckinna Jun 11, 2024
28c1791
refactor: use existing lemma
jamesmckinna Jun 14, 2024
d1f052c
Merge branch 'master' into solver-refactor
jamesmckinna Jun 19, 2024
7e8f2e8
fixed deprecations to `v2.2`
jamesmckinna Jul 26, 2024
f83f962
v2.2 `CHANGELOG`
jamesmckinna Jul 27, 2024
fbcb318
Merge branch 'master' into solver-refactor
jamesmckinna Jul 27, 2024
d75b2c9
use `zipWith`!
jamesmckinna Jul 28, 2024
b876988
remove dependency on `Data.Nat.GeneralisedArithmetic.fold`
jamesmckinna Jul 30, 2024
b7a819f
simplify `import` dependencies
jamesmckinna Jul 30, 2024
2940a78
refactor: tweaks
jamesmckinna Jul 30, 2024
74e6a9f
Merge branch 'master' into solver-refactor
jamesmckinna Aug 12, 2024
578691a
Merge branch 'master' into solver-refactor
jamesmckinna Aug 14, 2024
45dc3ca
Merge branch 'master' into solver-refactor
jamesmckinna Aug 14, 2024
4042438
Merge branch 'master' into solver-refactor
jamesmckinna Aug 24, 2024
ef864ee
change `Level` parameterisation of `NormalAPI`
jamesmckinna Aug 29, 2024
e35ae9f
clarify exports from `module R`?
jamesmckinna Aug 29, 2024
eb71dd0
deprecations
jamesmckinna Aug 29, 2024
f218b8c
knock-on effects
jamesmckinna Aug 29, 2024
1c02f94
deprecations
jamesmckinna Aug 29, 2024
ec98888
fixed exports of `Normal` form instantiations
jamesmckinna Aug 29, 2024
45379e6
fixed exports of `Solver`
jamesmckinna Aug 29, 2024
e1c3ab1
fixed deprecations in `CHANGELOG`
jamesmckinna Aug 29, 2024
07e005a
note on import strategy in `CHANGELOG`
jamesmckinna Aug 29, 2024
a0572d5
name of `Solver` module
jamesmckinna Aug 29, 2024
8714854
Merge branch 'master' into solver-refactor
MatthewDaggitt Sep 3, 2024
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Loading
Jump to
Jump to file
Failed to load files.
Loading
Diff view
Diff view
33 changes: 33 additions & 0 deletions CHANGELOG.md
Original file line number Diff line number Diff line change
Expand Up @@ -69,6 +69,22 @@ Deprecated names
1×-identityʳ ↦ ×-homo-1
```

* In `Algebra.Solver.CommutativeMonoid`:
```agda
normalise-correct ↦ Algebra.Solver.CommutativeMonoid.Normal.correct
```

* In `Algebra.Solver.IdempotentCommutativeMonoid`:
```agda
normalise-correct ↦ Algebra.Solver.IdempotentCommutativeMonoid.Normal.correct
```

* In `Algebra.Solver.Monoid`:
```agda
homomorphic ↦ Algebra.Solver.Monoid.Normal.comp-correct
normalise-correct ↦ Algebra.Solver.Monoid.Normal.correct
```

* In `Algebra.Structures.IsGroup`:
```agda
_-_ ↦ _//_
Expand Down Expand Up @@ -132,6 +148,17 @@ New modules
Algebra.Morphism.Construct.Terminal
```

* Refactoring of the `Algebra.Solver.*Monoid` implementations, via
a single `Tactic` module API based on the existing `Expr`, and
a common `Normal`-form API:
jamesmckinna marked this conversation as resolved.
Show resolved Hide resolved
```agda
Algebra.Solver.CommutativeMonoid.Normal
Algebra.Solver.IdempotentCommutativeMonoid.Normal
Algebra.Solver.Monoid.Expression
Algebra.Solver.Monoid.Normal
Algebra.Solver.Monoid.Tactic
```

* Pointwise and equality relations over indexed containers:
```agda
Data.Container.Indexed.Relation.Binary.Pointwise
Expand Down Expand Up @@ -374,6 +401,12 @@ Additions to existing modules
idem-×-homo-* : (_*_ IdempotentOn x) → (m × x) * (n × x) ≈ (m ℕ.* n) × x
```

* In `Algebra.Solver.Ring`
```agda
Env : ℕ → Set _
Env = Vec Carrier
```

* In `Algebra.Structures`
```agda
record IsSuccessorSet (suc# : Op₁ A) (zero# : A) : Set _
Expand Down
193 changes: 17 additions & 176 deletions src/Algebra/Solver/CommutativeMonoid.agda
Original file line number Diff line number Diff line change
Expand Up @@ -8,195 +8,36 @@

{-# OPTIONS --cubical-compatible --safe #-}

open import Algebra
open import Algebra.Bundles using (CommutativeMonoid)

module Algebra.Solver.CommutativeMonoid {m₁ m₂} (M : CommutativeMonoid m₁ m₂) where
module Algebra.Solver.CommutativeMonoid {c ℓ} (M : CommutativeMonoid c ℓ) where

open import Data.Fin.Base using (Fin; zero; suc)
open import Data.Maybe.Base as Maybe
using (Maybe; decToMaybe; From-just; from-just)
open import Data.Nat as ℕ using (ℕ; zero; suc; _+_)
open import Data.Nat.GeneralisedArithmetic using (fold)
open import Data.Product.Base using (_×_; uncurry)
open import Data.Vec.Base using (Vec; []; _∷_; lookup; replicate)
import Algebra.Solver.CommutativeMonoid.Normal as Normal
import Algebra.Solver.Monoid.Tactic as Tactic

open import Function.Base using (_∘_)

import Relation.Binary.Reasoning.Setoid as ≈-Reasoning
import Relation.Binary.Reflection as Reflection
import Relation.Nullary.Decidable as Dec
import Data.Vec.Relation.Binary.Pointwise.Inductive as Pointwise

open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_)
open import Relation.Nullary.Decidable using (Dec)

open CommutativeMonoid M
open ≈-Reasoning setoid

private
variable
n : ℕ

------------------------------------------------------------------------
-- Monoid expressions

-- There is one constructor for every operation, plus one for
-- variables; there may be at most n variables.

infixr 5 _⊕_
infixr 10 _•_

data Expr (n : ℕ) : Set where
var : Fin n → Expr n
id : Expr n
_⊕_ : Expr n → Expr n → Expr n

-- An environment contains one value for every variable.

Env : ℕ → Set _
Env n = Vec Carrier n

-- The semantics of an expression is a function from an environment to
-- a value.

⟦_⟧ : Expr n → Env n → Carrier
⟦ var x ⟧ ρ = lookup ρ x
⟦ id ⟧ ρ = ε
⟦ e₁ ⊕ e₂ ⟧ ρ = ⟦ e₁ ⟧ ρ ∙ ⟦ e₂ ⟧ ρ
open CommutativeMonoid M using (monoid)

------------------------------------------------------------------------
-- Normal forms

-- A normal form is a vector of multiplicities (a bag).

Normal : ℕ → Set
Normal n = Vec ℕ n

-- The semantics of a normal form.

⟦_⟧⇓ : Normal n → Env n → Carrier
⟦ [] ⟧⇓ _ = ε
⟦ n ∷ v ⟧⇓ (a ∷ ρ) = fold (⟦ v ⟧⇓ ρ) (a ∙_) n
open module N = Normal M public
renaming (correct to normalise-correct)

------------------------------------------------------------------------
-- Constructions on normal forms

-- The empty bag.

empty : Normal n
empty = replicate _ 0
-- Tactic

-- A singleton bag.
open Tactic monoid (record { N }) public

sg : (i : Fin n) → Normal n
sg zero = 1 ∷ empty
sg (suc i) = 0 ∷ sg i

-- The composition of normal forms.

_•_ : (v w : Normal n) → Normal n
[] • [] = []
(l ∷ v) • (m ∷ w) = l + m ∷ v • w

------------------------------------------------------------------------
-- Correctness of the constructions on normal forms

-- The empty bag stands for the unit ε.

empty-correct : (ρ : Env n) → ⟦ empty ⟧⇓ ρ ≈ ε
empty-correct [] = refl
empty-correct (a ∷ ρ) = empty-correct ρ

-- The singleton bag stands for a single variable.

sg-correct : (x : Fin n) (ρ : Env n) → ⟦ sg x ⟧⇓ ρ ≈ lookup ρ x
sg-correct zero (x ∷ ρ) = begin
x ∙ ⟦ empty ⟧⇓ ρ ≈⟨ ∙-congˡ (empty-correct ρ) ⟩
x ∙ ε ≈⟨ identityʳ _ ⟩
x ∎
sg-correct (suc x) (m ∷ ρ) = sg-correct x ρ

-- Normal form composition corresponds to the composition of the monoid.

comp-correct : (v w : Normal n) (ρ : Env n) →
⟦ v • w ⟧⇓ ρ ≈ (⟦ v ⟧⇓ ρ ∙ ⟦ w ⟧⇓ ρ)
comp-correct [] [] ρ = sym (identityˡ _)
comp-correct (l ∷ v) (m ∷ w) (a ∷ ρ) = lemma l m (comp-correct v w ρ)
where
flip12 : ∀ a b c → a ∙ (b ∙ c) ≈ b ∙ (a ∙ c)
flip12 a b c = begin
a ∙ (b ∙ c) ≈⟨ sym (assoc _ _ _) ⟩
(a ∙ b) ∙ c ≈⟨ ∙-congʳ (comm _ _) ⟩
(b ∙ a) ∙ c ≈⟨ assoc _ _ _ ⟩
b ∙ (a ∙ c) ∎
lemma : ∀ l m {d b c} (p : d ≈ b ∙ c) →
fold d (a ∙_) (l + m) ≈ fold b (a ∙_) l ∙ fold c (a ∙_) m
lemma zero zero p = p
lemma zero (suc m) p = trans (∙-congˡ (lemma zero m p)) (flip12 _ _ _)
lemma (suc l) m p = trans (∙-congˡ (lemma l m p)) (sym (assoc a _ _))

------------------------------------------------------------------------
-- Normalization

-- A normaliser.

normalise : Expr n → Normal n
normalise (var x) = sg x
normalise id = empty
normalise (e₁ ⊕ e₂) = normalise e₁ • normalise e₂

-- The normaliser preserves the semantics of the expression.

normalise-correct : (e : Expr n) (ρ : Env n) →
⟦ normalise e ⟧⇓ ρ ≈ ⟦ e ⟧ ρ
normalise-correct (var x) ρ = sg-correct x ρ
normalise-correct id ρ = empty-correct ρ
normalise-correct (e₁ ⊕ e₂) ρ = begin

⟦ normalise e₁ • normalise e₂ ⟧⇓ ρ

≈⟨ comp-correct (normalise e₁) (normalise e₂) ρ ⟩

⟦ normalise e₁ ⟧⇓ ρ ∙ ⟦ normalise e₂ ⟧⇓ ρ

≈⟨ ∙-cong (normalise-correct e₁ ρ) (normalise-correct e₂ ρ) ⟩

⟦ e₁ ⟧ ρ ∙ ⟦ e₂ ⟧ ρ

-- DEPRECATED NAMES
------------------------------------------------------------------------
-- "Tactic.

open module R = Reflection
setoid var ⟦_⟧ (⟦_⟧⇓ ∘ normalise) normalise-correct
public using (solve; _⊜_)

-- We can decide if two normal forms are /syntactically/ equal.

infix 5 _≟_

_≟_ : (nf₁ nf₂ : Normal n) → Dec (nf₁ ≡ nf₂)
nf₁ ≟ nf₂ = Dec.map Pointwise-≡↔≡ (decidable ℕ._≟_ nf₁ nf₂)
where open Pointwise

-- We can also give a sound, but not necessarily complete, procedure
-- for determining if two expressions have the same semantics.

prove′ : (e₁ e₂ : Expr n) → Maybe (∀ ρ → ⟦ e₁ ⟧ ρ ≈ ⟦ e₂ ⟧ ρ)
prove′ e₁ e₂ =
Maybe.map lemma (decToMaybe (normalise e₁ ≟ normalise e₂))
where
lemma : normalise e₁ ≡ normalise e₂ → ∀ ρ → ⟦ e₁ ⟧ ρ ≈ ⟦ e₂ ⟧ ρ
lemma eq ρ =
R.prove ρ e₁ e₂ (begin
⟦ normalise e₁ ⟧⇓ ρ ≡⟨ ≡.cong (λ e → ⟦ e ⟧⇓ ρ) eq ⟩
⟦ normalise e₂ ⟧⇓ ρ ∎)

-- This procedure can be combined with from-just.
-- Please use the new names as continuing support for the old names is
-- not guaranteed.

prove : ∀ n (e₁ e₂ : Expr n) → From-just (prove′ e₁ e₂)
prove _ e₁ e₂ = from-just (prove′ e₁ e₂)
-- Version 2.1

-- prove : ∀ n (es : Expr n × Expr n) →
-- From-just (uncurry prove′ es)
-- prove _ = from-just ∘ uncurry prove′
{-# WARNING_ON_USAGE normalise-correct
"Warning: normalise-correct was deprecated in v2.1.
Please use Algebra.Solver.CommutativeMonoid.Normal.correct instead."
#-}
Loading