{-# OPTIONS --cubical-compatible --safe #-}
open import Algebra
module Algebra.Solver.CommutativeMonoid {m₁ m₂} (M : CommutativeMonoid m₁ m₂) where
open import Data.Fin.Base using (Fin; zero; suc)
open import Data.Maybe.Base as Maybe
using (Maybe; 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)
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.Consequences using (dec⇒weaklyDec)
open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_)
open import Relation.Nullary.Decidable as Dec using (Dec)
open CommutativeMonoid M
open ≈-Reasoning setoid
private
variable
n : ℕ
infixr 5 _⊕_
infixr 10 _•_
data Expr (n : ℕ) : Set where
var : Fin n → Expr n
id : Expr n
_⊕_ : Expr n → Expr n → Expr n
Env : ℕ → Set _
Env n = Vec Carrier n
⟦_⟧ : Expr n → Env n → Carrier
⟦ var x ⟧ ρ = lookup ρ x
⟦ id ⟧ ρ = ε
⟦ e₁ ⊕ e₂ ⟧ ρ = ⟦ e₁ ⟧ ρ ∙ ⟦ e₂ ⟧ ρ
Normal : ℕ → Set
Normal n = Vec ℕ n
⟦_⟧⇓ : Normal n → Env n → Carrier
⟦ [] ⟧⇓ _ = ε
⟦ n ∷ v ⟧⇓ (a ∷ ρ) = fold (⟦ v ⟧⇓ ρ) (a ∙_) n
empty : Normal n
empty = replicate _ 0
sg : (i : Fin n) → Normal n
sg zero = 1 ∷ empty
sg (suc i) = 0 ∷ sg i
_•_ : (v w : Normal n) → Normal n
[] • [] = []
(l ∷ v) • (m ∷ w) = l + m ∷ v • w
empty-correct : (ρ : Env n) → ⟦ empty ⟧⇓ ρ ≈ ε
empty-correct [] = refl
empty-correct (a ∷ ρ) = empty-correct ρ
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 ρ
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 _ _))
normalise : Expr n → Normal n
normalise (var x) = sg x
normalise id = empty
normalise (e₁ ⊕ e₂) = normalise e₁ • normalise e₂
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₂ ⟧ ρ
∎
open module R = Reflection
setoid var ⟦_⟧ (⟦_⟧⇓ ∘ normalise) normalise-correct
public using (solve; _⊜_)
infix 5 _≟_
_≟_ : (nf₁ nf₂ : Normal n) → Dec (nf₁ ≡ nf₂)
nf₁ ≟ nf₂ = Dec.map Pointwise-≡↔≡ (decidable ℕ._≟_ nf₁ nf₂)
where open Pointwise
prove′ : (e₁ e₂ : Expr n) → Maybe (∀ ρ → ⟦ e₁ ⟧ ρ ≈ ⟦ e₂ ⟧ ρ)
prove′ e₁ e₂ =
Maybe.map lemma (dec⇒weaklyDec _≟_ (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₂ ⟧⇓ ρ ∎)
prove : ∀ n (e₁ e₂ : Expr n) → From-just (prove′ e₁ e₂)
prove _ e₁ e₂ = from-just (prove′ e₁ e₂)