154
|
1 module stack where
|
|
2
|
161
|
3 open import Relation.Binary.PropositionalEquality
|
477
|
4 open import Relation.Binary.Core
|
161
|
5
|
179
|
6 data Nat : Set where
|
|
7 zero : Nat
|
|
8 suc : Nat -> Nat
|
|
9
|
161
|
10 data Bool : Set where
|
|
11 True : Bool
|
|
12 False : Bool
|
164
|
13
|
477
|
14 -- equal : {a : Set} -> a -> a -> Bool
|
|
15 -- equal x y with (x ≡)y)
|
|
16 -- equal x .x | refl = True
|
|
17 -- equal _ _ | _ = False
|
|
18
|
|
19
|
161
|
20 data Maybe (a : Set) : Set where
|
|
21 Nothing : Maybe a
|
|
22 Just : a -> Maybe a
|
|
23
|
|
24 record Stack {a t : Set} (stackImpl : Set) : Set where
|
|
25 field
|
|
26 stack : stackImpl
|
|
27 push : stackImpl -> a -> (stackImpl -> t) -> t
|
|
28 pop : stackImpl -> (stackImpl -> Maybe a -> t) -> t
|
155
|
29
|
477
|
30 open Stack
|
427
|
31
|
477
|
32 pushStack : {a t si : Set} -> Stack si -> a -> (Stack si -> t) -> t
|
|
33 pushStack {a} {t} s0 d next = push s0 (stack s0) d (\s1 -> next (record s0 {stack = s1} ))
|
|
34
|
|
35 popStack : {a t si : Set} -> Stack si -> (Stack si -> Maybe a -> t) -> t
|
|
36 popStack {a} {t} s0 next = pop s0 (stack s0) (\s1 d1 -> next (record s0 {stack = s1}) d1 )
|
427
|
37
|
|
38
|
161
|
39 data Element (a : Set) : Set where
|
|
40 cons : a -> Maybe (Element a) -> Element a
|
|
41
|
|
42 datum : {a : Set} -> Element a -> a
|
|
43 datum (cons a _) = a
|
|
44
|
|
45 next : {a : Set} -> Element a -> Maybe (Element a)
|
|
46 next (cons _ n) = n
|
|
47
|
|
48
|
164
|
49 {-
|
|
50 -- cannot define recrusive record definition. so use linked list with maybe.
|
161
|
51 record Element {l : Level} (a : Set l) : Set (suc l) where
|
|
52 field
|
164
|
53 datum : a -- `data` is reserved by Agda.
|
161
|
54 next : Maybe (Element a)
|
|
55 -}
|
155
|
56
|
|
57
|
164
|
58
|
161
|
59 record SingleLinkedStack (a : Set) : Set where
|
|
60 field
|
|
61 top : Maybe (Element a)
|
|
62 open SingleLinkedStack
|
155
|
63
|
161
|
64 pushSingleLinkedStack : {Data t : Set} -> SingleLinkedStack Data -> Data -> (Code : SingleLinkedStack Data -> t) -> t
|
|
65 pushSingleLinkedStack stack datum next = next stack1
|
|
66 where
|
|
67 element = cons datum (top stack)
|
164
|
68 stack1 = record {top = Just element}
|
161
|
69
|
155
|
70
|
161
|
71 popSingleLinkedStack : {a t : Set} -> SingleLinkedStack a -> (Code : SingleLinkedStack a -> (Maybe a) -> t) -> t
|
|
72 popSingleLinkedStack stack cs with (top stack)
|
|
73 ... | Nothing = cs stack Nothing
|
|
74 ... | Just d = cs stack1 (Just data1)
|
154
|
75 where
|
161
|
76 data1 = datum d
|
|
77 stack1 = record { top = (next d) }
|
154
|
78
|
|
79
|
161
|
80 emptySingleLinkedStack : {a : Set} -> SingleLinkedStack a
|
|
81 emptySingleLinkedStack = record {top = Nothing}
|
|
82
|
164
|
83 createSingleLinkedStack : {a b : Set} -> Stack {a} {b} (SingleLinkedStack a)
|
161
|
84 createSingleLinkedStack = record { stack = emptySingleLinkedStack
|
|
85 ; push = pushSingleLinkedStack
|
|
86 ; pop = popSingleLinkedStack
|
|
87 }
|
|
88
|
156
|
89
|
|
90
|
161
|
91 test01 : {a : Set} -> SingleLinkedStack a -> Maybe a -> Bool
|
|
92 test01 stack _ with (top stack)
|
|
93 ... | (Just _) = True
|
|
94 ... | Nothing = False
|
|
95
|
156
|
96
|
161
|
97 test02 : {a : Set} -> SingleLinkedStack a -> Bool
|
|
98 test02 stack = (popSingleLinkedStack stack) test01
|
156
|
99
|
165
|
100 test03 : {a : Set} -> a -> Bool
|
|
101 test03 v = pushSingleLinkedStack emptySingleLinkedStack v test02
|
156
|
102
|
477
|
103 testStack01 : {a : Set} -> a -> Bool
|
|
104 testStack01 v = pushStack createSingleLinkedStack v (
|
|
105 \s -> popStack s (\s1 d1 -> True))
|
|
106
|
|
107
|
161
|
108
|
165
|
109 lemma : {A : Set} {a : A} -> test03 a ≡ False
|
158
|
110 lemma = refl
|
179
|
111
|
|
112 id : {A : Set} -> A -> A
|
|
113 id a = a
|
|
114
|
|
115
|
|
116 n-push : {A : Set} {a : A} -> Nat -> SingleLinkedStack A -> SingleLinkedStack A
|
|
117 n-push zero s = s
|
|
118 n-push {A} {a} (suc n) s = pushSingleLinkedStack (n-push {A} {a} n s) a (\s -> s)
|
|
119
|
|
120 n-pop : {A : Set} {a : A} -> Nat -> SingleLinkedStack A -> SingleLinkedStack A
|
|
121 n-pop zero s = s
|
|
122 n-pop {A} {a} (suc n) s = popSingleLinkedStack (n-pop {A} {a} n s) (\s _ -> s)
|
|
123
|
|
124 open ≡-Reasoning
|
|
125
|
|
126 push-pop-equiv : {A : Set} {a : A} (s : SingleLinkedStack A) -> popSingleLinkedStack (pushSingleLinkedStack s a (\s -> s)) (\s _ -> s) ≡ s
|
|
127 push-pop-equiv s = refl
|
|
128
|
|
129 push-and-n-pop : {A : Set} {a : A} (n : Nat) (s : SingleLinkedStack A) -> n-pop {A} {a} (suc n) (pushSingleLinkedStack s a id) ≡ n-pop {A} {a} n s
|
|
130 push-and-n-pop zero s = refl
|
|
131 push-and-n-pop {A} {a} (suc n) s = begin
|
|
132 n-pop (suc (suc n)) (pushSingleLinkedStack s a id)
|
|
133 ≡⟨ refl ⟩
|
|
134 popSingleLinkedStack (n-pop (suc n) (pushSingleLinkedStack s a id)) (\s _ -> s)
|
|
135 ≡⟨ cong (\s -> popSingleLinkedStack s (\s _ -> s)) (push-and-n-pop n s) ⟩
|
|
136 popSingleLinkedStack (n-pop n s) (\s _ -> s)
|
|
137 ≡⟨ refl ⟩
|
|
138 n-pop (suc n) s
|
|
139 ∎
|
|
140
|
|
141
|
|
142 n-push-pop-equiv : {A : Set} {a : A} (n : Nat) (s : SingleLinkedStack A) -> (n-pop {A} {a} n (n-push {A} {a} n s)) ≡ s
|
|
143 n-push-pop-equiv zero s = refl
|
|
144 n-push-pop-equiv {A} {a} (suc n) s = begin
|
|
145 n-pop (suc n) (n-push (suc n) s)
|
|
146 ≡⟨ refl ⟩
|
|
147 n-pop (suc n) (pushSingleLinkedStack (n-push n s) a (\s -> s))
|
|
148 ≡⟨ push-and-n-pop n (n-push n s) ⟩
|
|
149 n-pop n (n-push n s)
|
180
|
150 ≡⟨ n-push-pop-equiv n s ⟩
|
179
|
151 s
|
|
152 ∎
|
181
|
153
|
|
154
|
|
155 n-push-pop-equiv-empty : {A : Set} {a : A} -> (n : Nat) -> n-pop {A} {a} n (n-push {A} {a} n emptySingleLinkedStack) ≡ emptySingleLinkedStack
|
|
156 n-push-pop-equiv-empty n = n-push-pop-equiv n emptySingleLinkedStack
|