154
|
1 module stack where
|
|
2
|
161
|
3 open import Relation.Binary.PropositionalEquality
|
|
4
|
|
5 data Bool : Set where
|
|
6 True : Bool
|
|
7 False : Bool
|
164
|
8
|
161
|
9 data Maybe (a : Set) : Set where
|
|
10 Nothing : Maybe a
|
|
11 Just : a -> Maybe a
|
|
12
|
|
13 record Stack {a t : Set} (stackImpl : Set) : Set where
|
|
14 field
|
|
15 stack : stackImpl
|
|
16 push : stackImpl -> a -> (stackImpl -> t) -> t
|
|
17 pop : stackImpl -> (stackImpl -> Maybe a -> t) -> t
|
155
|
18
|
161
|
19 data Element (a : Set) : Set where
|
|
20 cons : a -> Maybe (Element a) -> Element a
|
|
21
|
|
22 datum : {a : Set} -> Element a -> a
|
|
23 datum (cons a _) = a
|
|
24
|
|
25 next : {a : Set} -> Element a -> Maybe (Element a)
|
|
26 next (cons _ n) = n
|
|
27
|
|
28
|
164
|
29 {-
|
|
30 -- cannot define recrusive record definition. so use linked list with maybe.
|
161
|
31 record Element {l : Level} (a : Set l) : Set (suc l) where
|
|
32 field
|
164
|
33 datum : a -- `data` is reserved by Agda.
|
161
|
34 next : Maybe (Element a)
|
|
35 -}
|
155
|
36
|
|
37
|
164
|
38
|
161
|
39 record SingleLinkedStack (a : Set) : Set where
|
|
40 field
|
|
41 top : Maybe (Element a)
|
|
42 open SingleLinkedStack
|
155
|
43
|
161
|
44 pushSingleLinkedStack : {Data t : Set} -> SingleLinkedStack Data -> Data -> (Code : SingleLinkedStack Data -> t) -> t
|
|
45 pushSingleLinkedStack stack datum next = next stack1
|
|
46 where
|
|
47 element = cons datum (top stack)
|
164
|
48 stack1 = record {top = Just element}
|
161
|
49
|
155
|
50
|
161
|
51 popSingleLinkedStack : {a t : Set} -> SingleLinkedStack a -> (Code : SingleLinkedStack a -> (Maybe a) -> t) -> t
|
|
52 popSingleLinkedStack stack cs with (top stack)
|
|
53 ... | Nothing = cs stack Nothing
|
|
54 ... | Just d = cs stack1 (Just data1)
|
154
|
55 where
|
161
|
56 data1 = datum d
|
|
57 stack1 = record { top = (next d) }
|
154
|
58
|
|
59
|
161
|
60 emptySingleLinkedStack : {a : Set} -> SingleLinkedStack a
|
|
61 emptySingleLinkedStack = record {top = Nothing}
|
|
62
|
164
|
63 createSingleLinkedStack : {a b : Set} -> Stack {a} {b} (SingleLinkedStack a)
|
161
|
64 createSingleLinkedStack = record { stack = emptySingleLinkedStack
|
|
65 ; push = pushSingleLinkedStack
|
|
66 ; pop = popSingleLinkedStack
|
|
67 }
|
|
68
|
156
|
69
|
|
70
|
161
|
71 test01 : {a : Set} -> SingleLinkedStack a -> Maybe a -> Bool
|
|
72 test01 stack _ with (top stack)
|
|
73 ... | (Just _) = True
|
|
74 ... | Nothing = False
|
|
75
|
156
|
76
|
161
|
77 test02 : {a : Set} -> SingleLinkedStack a -> Bool
|
|
78 test02 stack = (popSingleLinkedStack stack) test01
|
156
|
79
|
165
|
80 test03 : {a : Set} -> a -> Bool
|
|
81 test03 v = pushSingleLinkedStack emptySingleLinkedStack v test02
|
156
|
82
|
161
|
83
|
165
|
84 lemma : {A : Set} {a : A} -> test03 a ≡ False
|
158
|
85 lemma = refl
|