forked from agda/agda-stdlib
-
Notifications
You must be signed in to change notification settings - Fork 0
/
Copy pathMonoid.agda
138 lines (104 loc) · 4.77 KB
/
Monoid.agda
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
60
61
62
63
64
65
66
67
68
69
70
71
72
73
74
75
76
77
78
79
80
81
82
83
84
85
86
87
88
89
90
91
92
93
94
95
96
97
98
99
100
101
102
103
104
105
106
107
108
109
110
111
112
113
114
115
116
117
118
119
120
121
122
123
124
125
126
127
128
129
130
131
132
133
134
135
136
137
138
------------------------------------------------------------------------
-- The Agda standard library
--
-- A solver for equations over monoids
------------------------------------------------------------------------
{-# OPTIONS --cubical-compatible --safe #-}
open import Algebra
module Algebra.Solver.Monoid {m₁ m₂} (M : Monoid m₁ m₂) where
open import Data.Fin.Base as Fin
import Data.Fin.Properties as Fin
open import Data.List.Base hiding (lookup)
import Data.List.Relation.Binary.Equality.DecPropositional as ListEq
open import Data.Maybe.Base as Maybe
using (Maybe; decToMaybe; From-just; from-just)
open import Data.Nat.Base using (ℕ)
open import Data.Product
open import Data.Vec.Base using (Vec; lookup)
open import Function.Base using (_∘_; _$_)
open import Relation.Binary.Definitions using (Decidable)
open import Relation.Binary.PropositionalEquality as P using (_≡_)
import Relation.Binary.Reflection
open import Relation.Nullary
import Relation.Nullary.Decidable as Dec
open Monoid M
open import Relation.Binary.Reasoning.Setoid setoid
------------------------------------------------------------------------
-- Monoid expressions
-- There is one constructor for every operation, plus one for
-- variables; there may be at most n variables.
infixr 5 _⊕_
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.
⟦_⟧ : ∀ {n} → Expr n → Env n → Carrier
⟦ var x ⟧ ρ = lookup ρ x
⟦ id ⟧ ρ = ε
⟦ e₁ ⊕ e₂ ⟧ ρ = ⟦ e₁ ⟧ ρ ∙ ⟦ e₂ ⟧ ρ
------------------------------------------------------------------------
-- Normal forms
-- A normal form is a list of variables.
Normal : ℕ → Set
Normal n = List (Fin n)
-- The semantics of a normal form.
⟦_⟧⇓ : ∀ {n} → Normal n → Env n → Carrier
⟦ [] ⟧⇓ ρ = ε
⟦ x ∷ nf ⟧⇓ ρ = lookup ρ x ∙ ⟦ nf ⟧⇓ ρ
-- A normaliser.
normalise : ∀ {n} → Expr n → Normal n
normalise (var x) = x ∷ []
normalise id = []
normalise (e₁ ⊕ e₂) = normalise e₁ ++ normalise e₂
-- The normaliser is homomorphic with respect to _++_/_∙_.
homomorphic : ∀ {n} (nf₁ nf₂ : Normal n) (ρ : Env n) →
⟦ nf₁ ++ nf₂ ⟧⇓ ρ ≈ (⟦ nf₁ ⟧⇓ ρ ∙ ⟦ nf₂ ⟧⇓ ρ)
homomorphic [] nf₂ ρ = begin
⟦ nf₂ ⟧⇓ ρ ≈⟨ sym $ identityˡ _ ⟩
ε ∙ ⟦ nf₂ ⟧⇓ ρ ∎
homomorphic (x ∷ nf₁) nf₂ ρ = begin
lookup ρ x ∙ ⟦ nf₁ ++ nf₂ ⟧⇓ ρ ≈⟨ ∙-congˡ (homomorphic nf₁ nf₂ ρ) ⟩
lookup ρ x ∙ (⟦ nf₁ ⟧⇓ ρ ∙ ⟦ nf₂ ⟧⇓ ρ) ≈⟨ sym $ assoc _ _ _ ⟩
lookup ρ x ∙ ⟦ nf₁ ⟧⇓ ρ ∙ ⟦ nf₂ ⟧⇓ ρ ∎
-- The normaliser preserves the semantics of the expression.
normalise-correct :
∀ {n} (e : Expr n) (ρ : Env n) → ⟦ normalise e ⟧⇓ ρ ≈ ⟦ e ⟧ ρ
normalise-correct (var x) ρ = begin
lookup ρ x ∙ ε ≈⟨ identityʳ _ ⟩
lookup ρ x ∎
normalise-correct id ρ = begin
ε ∎
normalise-correct (e₁ ⊕ e₂) ρ = begin
⟦ normalise e₁ ++ normalise e₂ ⟧⇓ ρ ≈⟨ homomorphic (normalise e₁) (normalise e₂) ρ ⟩
⟦ normalise e₁ ⟧⇓ ρ ∙ ⟦ normalise e₂ ⟧⇓ ρ ≈⟨ ∙-cong (normalise-correct e₁ ρ) (normalise-correct e₂ ρ) ⟩
⟦ e₁ ⟧ ρ ∙ ⟦ e₂ ⟧ ρ ∎
------------------------------------------------------------------------
-- "Tactic.
open module R = Relation.Binary.Reflection
setoid var ⟦_⟧ (⟦_⟧⇓ ∘ normalise) normalise-correct
public using (solve; _⊜_)
-- We can decide if two normal forms are /syntactically/ equal.
infix 5 _≟_
_≟_ : ∀ {n} → Decidable {A = Normal n} _≡_
nf₁ ≟ nf₂ = Dec.map′ ≋⇒≡ ≡⇒≋ (nf₁ ≋? nf₂)
where open ListEq Fin._≟_
-- We can also give a sound, but not necessarily complete, procedure
-- for determining if two expressions have the same semantics.
prove′ : ∀ {n} (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₁ ⟧⇓ ρ ≡⟨ P.cong (λ e → ⟦ e ⟧⇓ ρ) eq ⟩
⟦ normalise e₂ ⟧⇓ ρ ∎)
-- This procedure can be combined with from-just.
prove : ∀ n (es : Expr n × Expr n) →
From-just (uncurry prove′ es)
prove _ = from-just ∘ uncurry prove′