|
8 | 8 |
|
9 | 9 | {-# OPTIONS --cubical-compatible --safe #-} |
10 | 10 |
|
11 | | -open import Algebra |
| 11 | +open import Algebra.Bundles using (CommutativeMonoid) |
12 | 12 |
|
13 | | -module Algebra.Solver.CommutativeMonoid {m₁ m₂} (M : CommutativeMonoid m₁ m₂) where |
| 13 | +module Algebra.Solver.CommutativeMonoid {c ℓ} (M : CommutativeMonoid c ℓ) where |
14 | 14 |
|
15 | | -open import Data.Fin.Base using (Fin; zero; suc) |
16 | | -open import Data.Maybe.Base as Maybe |
17 | | - using (Maybe; From-just; from-just) |
18 | | -open import Data.Nat as ℕ using (ℕ; zero; suc; _+_) |
19 | | -open import Data.Nat.GeneralisedArithmetic using (fold) |
20 | | -open import Data.Product.Base using (_×_; uncurry) |
21 | | -open import Data.Vec.Base using (Vec; []; _∷_; lookup; replicate) |
| 15 | +import Algebra.Solver.CommutativeMonoid.Normal as Normal |
| 16 | +import Algebra.Solver.Monoid.Solver as Solver |
22 | 17 |
|
23 | | -open import Function.Base using (_∘_) |
24 | | - |
25 | | -import Relation.Binary.Reasoning.Setoid as ≈-Reasoning |
26 | | -import Relation.Binary.Reflection as Reflection |
27 | | -import Relation.Nullary.Decidable as Dec |
28 | | -import Data.Vec.Relation.Binary.Pointwise.Inductive as Pointwise |
29 | | - |
30 | | -open import Relation.Binary.Consequences using (dec⇒weaklyDec) |
31 | | -open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) |
32 | | -open import Relation.Nullary.Decidable as Dec using (Dec) |
33 | | - |
34 | | -open CommutativeMonoid M |
35 | | -open ≈-Reasoning setoid |
| 18 | +open CommutativeMonoid M using (monoid) |
36 | 19 |
|
37 | 20 | private |
38 | | - variable |
39 | | - n : ℕ |
40 | | - |
41 | | ------------------------------------------------------------------------- |
42 | | --- Monoid expressions |
43 | | - |
44 | | --- There is one constructor for every operation, plus one for |
45 | | --- variables; there may be at most n variables. |
46 | | - |
47 | | -infixr 5 _⊕_ |
48 | | -infixr 10 _•_ |
49 | | - |
50 | | -data Expr (n : ℕ) : Set where |
51 | | - var : Fin n → Expr n |
52 | | - id : Expr n |
53 | | - _⊕_ : Expr n → Expr n → Expr n |
54 | | - |
55 | | --- An environment contains one value for every variable. |
56 | | - |
57 | | -Env : ℕ → Set _ |
58 | | -Env n = Vec Carrier n |
59 | | - |
60 | | --- The semantics of an expression is a function from an environment to |
61 | | --- a value. |
| 21 | + module N = Normal M |
62 | 22 |
|
63 | | -⟦_⟧ : Expr n → Env n → Carrier |
64 | | -⟦ var x ⟧ ρ = lookup ρ x |
65 | | -⟦ id ⟧ ρ = ε |
66 | | -⟦ e₁ ⊕ e₂ ⟧ ρ = ⟦ e₁ ⟧ ρ ∙ ⟦ e₂ ⟧ ρ |
67 | 23 |
|
68 | 24 | ------------------------------------------------------------------------ |
69 | 25 | -- Normal forms |
70 | 26 |
|
71 | | --- A normal form is a vector of multiplicities (a bag). |
72 | | - |
73 | | -Normal : ℕ → Set |
74 | | -Normal n = Vec ℕ n |
75 | | - |
76 | | --- The semantics of a normal form. |
77 | | - |
78 | | -⟦_⟧⇓ : Normal n → Env n → Carrier |
79 | | -⟦ [] ⟧⇓ _ = ε |
80 | | -⟦ n ∷ v ⟧⇓ (a ∷ ρ) = fold (⟦ v ⟧⇓ ρ) (a ∙_) n |
| 27 | +open N public |
| 28 | + renaming (correct to normalise-correct) |
81 | 29 |
|
82 | 30 | ------------------------------------------------------------------------ |
83 | | --- Constructions on normal forms |
84 | | - |
85 | | --- The empty bag. |
86 | | - |
87 | | -empty : Normal n |
88 | | -empty = replicate _ 0 |
89 | | - |
90 | | --- A singleton bag. |
| 31 | +-- Proof procedures |
91 | 32 |
|
92 | | -sg : (i : Fin n) → Normal n |
93 | | -sg zero = 1 ∷ empty |
94 | | -sg (suc i) = 0 ∷ sg i |
| 33 | +open Solver monoid (record { N }) public |
95 | 34 |
|
96 | | --- The composition of normal forms. |
97 | | - |
98 | | -_•_ : (v w : Normal n) → Normal n |
99 | | -[] • [] = [] |
100 | | -(l ∷ v) • (m ∷ w) = l + m ∷ v • w |
101 | 35 |
|
102 | 36 | ------------------------------------------------------------------------ |
103 | | --- Correctness of the constructions on normal forms |
104 | | - |
105 | | --- The empty bag stands for the unit ε. |
106 | | - |
107 | | -empty-correct : (ρ : Env n) → ⟦ empty ⟧⇓ ρ ≈ ε |
108 | | -empty-correct [] = refl |
109 | | -empty-correct (a ∷ ρ) = empty-correct ρ |
110 | | - |
111 | | --- The singleton bag stands for a single variable. |
112 | | - |
113 | | -sg-correct : (x : Fin n) (ρ : Env n) → ⟦ sg x ⟧⇓ ρ ≈ lookup ρ x |
114 | | -sg-correct zero (x ∷ ρ) = begin |
115 | | - x ∙ ⟦ empty ⟧⇓ ρ ≈⟨ ∙-congˡ (empty-correct ρ) ⟩ |
116 | | - x ∙ ε ≈⟨ identityʳ _ ⟩ |
117 | | - x ∎ |
118 | | -sg-correct (suc x) (m ∷ ρ) = sg-correct x ρ |
119 | | - |
120 | | --- Normal form composition corresponds to the composition of the monoid. |
121 | | - |
122 | | -comp-correct : (v w : Normal n) (ρ : Env n) → |
123 | | - ⟦ v • w ⟧⇓ ρ ≈ (⟦ v ⟧⇓ ρ ∙ ⟦ w ⟧⇓ ρ) |
124 | | -comp-correct [] [] ρ = sym (identityˡ _) |
125 | | -comp-correct (l ∷ v) (m ∷ w) (a ∷ ρ) = lemma l m (comp-correct v w ρ) |
126 | | - where |
127 | | - flip12 : ∀ a b c → a ∙ (b ∙ c) ≈ b ∙ (a ∙ c) |
128 | | - flip12 a b c = begin |
129 | | - a ∙ (b ∙ c) ≈⟨ sym (assoc _ _ _) ⟩ |
130 | | - (a ∙ b) ∙ c ≈⟨ ∙-congʳ (comm _ _) ⟩ |
131 | | - (b ∙ a) ∙ c ≈⟨ assoc _ _ _ ⟩ |
132 | | - b ∙ (a ∙ c) ∎ |
133 | | - lemma : ∀ l m {d b c} (p : d ≈ b ∙ c) → |
134 | | - fold d (a ∙_) (l + m) ≈ fold b (a ∙_) l ∙ fold c (a ∙_) m |
135 | | - lemma zero zero p = p |
136 | | - lemma zero (suc m) p = trans (∙-congˡ (lemma zero m p)) (flip12 _ _ _) |
137 | | - lemma (suc l) m p = trans (∙-congˡ (lemma l m p)) (sym (assoc a _ _)) |
138 | | - |
| 37 | +-- DEPRECATED NAMES |
139 | 38 | ------------------------------------------------------------------------ |
140 | | --- Normalization |
141 | | - |
142 | | --- A normaliser. |
143 | | - |
144 | | -normalise : Expr n → Normal n |
145 | | -normalise (var x) = sg x |
146 | | -normalise id = empty |
147 | | -normalise (e₁ ⊕ e₂) = normalise e₁ • normalise e₂ |
148 | | - |
149 | | --- The normaliser preserves the semantics of the expression. |
150 | | - |
151 | | -normalise-correct : (e : Expr n) (ρ : Env n) → |
152 | | - ⟦ normalise e ⟧⇓ ρ ≈ ⟦ e ⟧ ρ |
153 | | -normalise-correct (var x) ρ = sg-correct x ρ |
154 | | -normalise-correct id ρ = empty-correct ρ |
155 | | -normalise-correct (e₁ ⊕ e₂) ρ = begin |
156 | | - |
157 | | - ⟦ normalise e₁ • normalise e₂ ⟧⇓ ρ |
158 | | - |
159 | | - ≈⟨ comp-correct (normalise e₁) (normalise e₂) ρ ⟩ |
160 | | - |
161 | | - ⟦ normalise e₁ ⟧⇓ ρ ∙ ⟦ normalise e₂ ⟧⇓ ρ |
162 | | - |
163 | | - ≈⟨ ∙-cong (normalise-correct e₁ ρ) (normalise-correct e₂ ρ) ⟩ |
164 | | - |
165 | | - ⟦ e₁ ⟧ ρ ∙ ⟦ e₂ ⟧ ρ |
166 | | - ∎ |
167 | | - |
168 | | ------------------------------------------------------------------------- |
169 | | --- "Tactic. |
170 | | - |
171 | | -open module R = Reflection |
172 | | - setoid var ⟦_⟧ (⟦_⟧⇓ ∘ normalise) normalise-correct |
173 | | - public using (solve; _⊜_) |
174 | | - |
175 | | --- We can decide if two normal forms are /syntactically/ equal. |
176 | | - |
177 | | -infix 5 _≟_ |
178 | | - |
179 | | -_≟_ : (nf₁ nf₂ : Normal n) → Dec (nf₁ ≡ nf₂) |
180 | | -nf₁ ≟ nf₂ = Dec.map Pointwise-≡↔≡ (decidable ℕ._≟_ nf₁ nf₂) |
181 | | - where open Pointwise |
182 | | - |
183 | | --- We can also give a sound, but not necessarily complete, procedure |
184 | | --- for determining if two expressions have the same semantics. |
185 | | - |
186 | | -prove′ : (e₁ e₂ : Expr n) → Maybe (∀ ρ → ⟦ e₁ ⟧ ρ ≈ ⟦ e₂ ⟧ ρ) |
187 | | -prove′ e₁ e₂ = |
188 | | - Maybe.map lemma (dec⇒weaklyDec _≟_ (normalise e₁) (normalise e₂)) |
189 | | - where |
190 | | - lemma : normalise e₁ ≡ normalise e₂ → ∀ ρ → ⟦ e₁ ⟧ ρ ≈ ⟦ e₂ ⟧ ρ |
191 | | - lemma eq ρ = |
192 | | - R.prove ρ e₁ e₂ (begin |
193 | | - ⟦ normalise e₁ ⟧⇓ ρ ≡⟨ ≡.cong (λ e → ⟦ e ⟧⇓ ρ) eq ⟩ |
194 | | - ⟦ normalise e₂ ⟧⇓ ρ ∎) |
195 | | - |
196 | | --- This procedure can be combined with from-just. |
| 39 | +-- Please use the new names as continuing support for the old names is |
| 40 | +-- not guaranteed. |
197 | 41 |
|
198 | | -prove : ∀ n (e₁ e₂ : Expr n) → From-just (prove′ e₁ e₂) |
199 | | -prove _ e₁ e₂ = from-just (prove′ e₁ e₂) |
| 42 | +-- Version 2.2 |
200 | 43 |
|
201 | | --- prove : ∀ n (es : Expr n × Expr n) → |
202 | | --- From-just (uncurry prove′ es) |
203 | | --- prove _ = from-just ∘ uncurry prove′ |
| 44 | +{-# WARNING_ON_USAGE normalise-correct |
| 45 | +"Warning: normalise-correct was deprecated in v2.2. |
| 46 | +Please use Algebra.Solver.CommutativeMonoid.Normal.correct instead." |
| 47 | +#-} |
0 commit comments