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