-
Notifications
You must be signed in to change notification settings - Fork 3
/
GlobalState.agda
82 lines (67 loc) · 1.94 KB
/
GlobalState.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
{-# OPTIONS --rewriting #-}
module Examples.Decalf.GlobalState where
open import Algebra.Cost
costMonoid = ℕ-CostMonoid
open CostMonoid costMonoid using (ℂ)
open import Calf costMonoid
open import Calf.Data.Nat as Nat using (nat; _*_)
open import Calf.Data.Equality as Eq using (_≡_; module ≡-Reasoning)
open import Function
S : tp⁺
S = nat
variable
s s₁ s₂ : val S
postulate
get : (X : tp⁻) → (val S → cmp X) → cmp X
set : (X : tp⁻) → val S → cmp X → cmp X
get/get : {e : val S → val S → cmp X} →
(get X λ s₁ → get X λ s₂ → e s₁ s₂) ≡ (get X λ s → e s s)
get/set : {e : cmp X} →
(get X λ s → set X s e) ≡ e
set/get : {e : val S → cmp X} →
set X s (get X e) ≡ set X s (e s)
set/set : {e : cmp X} →
set X s₁ (set X s₂ e) ≡ set X s₂ e
get/step : (c : ℂ) {e : val S → cmp X} →
step X c (get X e) ≡ get X λ s → step X c (e s)
set/step : (c : ℂ) {e : cmp X} →
step X c (set X s e) ≡ set X s (step X c e)
module StateDependentCost where
open import Examples.Decalf.Basic
e : cmp (F nat)
e =
get (F nat) λ n →
bind (F nat) (double n) λ n' →
set (F nat) n' $
ret n
e/bound : cmp (F nat)
e/bound =
get (F nat) λ n →
set (F nat) (2 * n) $
step (F nat) n $
ret n
e/has-cost : e ≡ e/bound
e/has-cost =
Eq.cong (get (F nat)) $ funext λ n →
let open ≡-Reasoning in
begin
( bind (F nat) (double n) λ n' →
set (F nat) n' $
ret n
)
≡⟨ Eq.cong (λ e₁ → bind (F nat) e₁ λ n' → set (F nat) n' (ret n)) (double/has-cost n) ⟩
( bind (F nat) (step (F nat) n (ret (2 * n))) λ n' →
set (F nat) n' $
ret n
)
≡⟨⟩
( step (F nat) n $
set (F nat) (2 * n) $
ret n
)
≡⟨ set/step n ⟩
( set (F nat) (2 * n) $
step (F nat) n $
ret n
)
∎